| author | wenzelm | 
| Sun, 29 Sep 2024 21:16:17 +0200 | |
| changeset 81008 | d0cd220d8e8b | 
| parent 80914 | d97fdabd9e2b | 
| child 82520 | 1b17f0a41aa3 | 
| permissions | -rw-r--r-- | 
| 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: 
78250diff
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: 
78248diff
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 | 18 | lemma approachable_lt_le2: | 
| 78685 | 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)" | 
| 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 | 51 |     with 1 obtain U where U: "open U" "x \<in> U" "\<forall>y\<in>U. indicator A y \<in> ({0<..<2} :: real set)"
 | 
| 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 | 61 |     with 1 obtain U where U: "open U" "x \<in> U" "\<forall>y\<in>U. indicator A y \<in> ({-1<..<1} :: real set)"
 | 
| 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 | 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 | 89 | lemma islimpt_closure: | 
| 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" | |
| 91 | using closure_def by fastforce | |
| 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: 
69753diff
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 | 95 | proof - | 
| 96 | have "\<And>U x. \<lbrakk>closed U; S = T \<inter> U; x islimpt S; x \<in> T\<rbrakk> \<Longrightarrow> x \<in> S" | |
| 97 | by (meson IntI closed_limpt inf_le2 islimpt_subset) | |
| 98 | then show ?thesis | |
| 99 | by (metis closed_closure closedin_closed closedin_imp_subset islimpt_closure) | |
| 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: 
69753diff
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 | 110 | text \<open>If a connnected set is written as the union of two nonempty closed sets, | 
| 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> {}"
 | 
| 78685 | 116 | by (metis assms closed_Un connected_closed_set) | 
| 69616 
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: 
69753diff
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: 
69753diff
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: 
69753diff
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: 
69753diff
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: | 
| 78685 | 129 | "\<lbrakk>compact S; closedin (top_of_set S) T\<rbrakk> \<Longrightarrow> compact T" | 
| 130 | by (metis closedin_closed compact_Int_closed) | |
| 69616 
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 | 134 | shows "compact S \<Longrightarrow> (closedin (top_of_set S) T \<longleftrightarrow> compact T \<and> T \<subseteq> S)" | 
| 78685 | 135 | by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) | 
| 69616 
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: 
69753diff
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: 
69753diff
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: 
69753diff
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: 
69753diff
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: 
69753diff
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 | 170 | lemma closure_Int_ballI: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 171 |   assumes "\<And>U. \<lbrakk>openin (top_of_set S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
 | 
| 69622 | 172 | shows "S \<subseteq> closure T" | 
| 173 | proof (clarsimp simp: closure_iff_nhds_not_empty) | |
| 174 | fix x and A and V | |
| 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: 
69753diff
changeset | 176 | then have "openin (top_of_set S) (A \<inter> V \<inter> S)" | 
| 76894 | 177 | by (simp add: inf_absorb2 openin_subtopology_Int) | 
| 69622 | 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>
 | 
| 179 | by auto | |
| 76894 | 180 | ultimately show False | 
| 181 |     using \<open>T \<inter> A = {}\<close> assms by fastforce
 | |
| 69622 | 182 | qed | 
| 183 | ||
| 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: 
69753diff
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: 
69753diff
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: 
69753diff
changeset | 189 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
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: 
69753diff
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: 
69753diff
changeset | 192 | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 193 | lemma connected_Int_frontier: | 
| 76894 | 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 | 196 | apply (drule_tac x="S \<inter> interior T" in spec, safe) | 
| 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 | 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: 
69753diff
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: 
69753diff
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: 
69753diff
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: 
69753diff
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 | 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 | 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: 
69753diff
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 | |
| 78685 | 307 | lemma continuous_image_closure_subset: | 
| 308 | assumes "continuous_on A f" "closure B \<subseteq> A" | |
| 309 | shows "f ` closure B \<subseteq> closure (f ` B)" | |
| 310 | using assms by (meson closed_closure closure_subset continuous_on_subset image_closure_subset) | |
| 311 | ||
| 70136 | 312 | 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 | 313 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78685diff
changeset | 314 | definition constant_on (infixl \<open>(constant'_on)\<close> 50) | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 315 | 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 | 316 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 317 | 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 | 318 | unfolding constant_on_def by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 319 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 320 | lemma injective_not_constant: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 321 |   fixes S :: "'a::{perfect_space} set"
 | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 322 |   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: 
78250diff
changeset | 323 | unfolding constant_on_def | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 324 | 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 | 325 | |
| 77138 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76894diff
changeset | 326 | 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: 
76894diff
changeset | 327 | 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: 
76894diff
changeset | 328 | 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: 
76894diff
changeset | 329 | 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: 
76894diff
changeset | 330 | |
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76894diff
changeset | 331 | 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: 
76894diff
changeset | 332 | "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: 
76894diff
changeset | 333 | 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: 
76894diff
changeset | 334 | |
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76894diff
changeset | 335 | 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: 
76894diff
changeset | 336 | 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: 
76894diff
changeset | 337 | 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: 
76894diff
changeset | 338 | 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: 
76894diff
changeset | 339 | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 340 | lemma constant_on_closureI: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 341 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 342 | 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: 
76894diff
changeset | 343 | 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: 
76894diff
changeset | 344 | 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: 
76894diff
changeset | 345 | by metis | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 346 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 347 | |
| 70136 | 348 | 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 | 349 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 350 | lemma continuous_on_Un_local: | 
| 76894 | 351 | "\<lbrakk>closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T; | 
| 352 | continuous_on S f; continuous_on T f\<rbrakk> | |
| 353 | \<Longrightarrow> continuous_on (S \<union> T) f" | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 354 | unfolding continuous_on closedin_limpt | 
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 355 | 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 | 356 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 357 | lemma continuous_on_cases_local: | 
| 76894 | 358 | "\<lbrakk>closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T; | 
| 359 | continuous_on S f; continuous_on T g; | |
| 360 | \<And>x. \<lbrakk>x \<in> S \<and> \<not>P x \<or> x \<in> T \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk> | |
| 361 | \<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 | 362 | 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 | 363 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 364 | lemma continuous_on_cases_le: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 365 | fixes h :: "'a :: topological_space \<Rightarrow> real" | 
| 76894 | 366 |   assumes "continuous_on {x \<in> S. h x \<le> a} f"
 | 
| 367 |       and "continuous_on {x \<in> S. a \<le> h x} g"
 | |
| 368 | and h: "continuous_on S h" | |
| 369 | and "\<And>x. \<lbrakk>x \<in> S; h x = a\<rbrakk> \<Longrightarrow> f x = g x" | |
| 370 | 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 | 371 | proof - | 
| 76894 | 372 | 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 | 373 | by force | 
| 76894 | 374 | 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 | 375 | by (rule continuous_closedin_preimage [OF h closed_atMost]) | 
| 76894 | 376 | 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 | 377 | by (rule continuous_closedin_preimage [OF h closed_atLeast]) | 
| 76894 | 378 |   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 | 379 | by auto | 
| 76894 | 380 |   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)"
 | 
| 381 | by (intro continuous_on_cases_local) (use 1 2 S assms in auto) | |
| 382 | then show ?thesis | |
| 383 | using S by force | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 384 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 385 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 386 | lemma continuous_on_cases_1: | 
| 76894 | 387 | fixes S :: "real set" | 
| 388 |   assumes "continuous_on {t \<in> S. t \<le> a} f"
 | |
| 389 |       and "continuous_on {t \<in> S. a \<le> t} g"
 | |
| 390 | and "a \<in> S \<Longrightarrow> f a = g a" | |
| 391 | 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 | 392 | using assms | 
| 71172 | 393 | 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 | 394 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 395 | |
| 70136 | 396 | 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 | 397 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 398 | lemma continuous_on_inverse_open_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 399 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 400 | and imf: "f ` S = T" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 401 | 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: 
69753diff
changeset | 402 | 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 | 403 | shows "continuous_on T g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 404 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 405 | from imf injf have gTS: "g ` T = S" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 406 | by force | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 407 | 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 | 408 | by force | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 409 | show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 410 | 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 | 411 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 412 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 413 | lemma continuous_on_inverse_closed_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 414 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 415 | and imf: "f ` S = T" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 416 | 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: 
69753diff
changeset | 417 | 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 | 418 | shows "continuous_on T g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 419 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 420 | from imf injf have gTS: "g ` T = S" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 421 | by force | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 422 | 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 | 423 | by force | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 424 | show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 425 | 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 | 426 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 427 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 428 | lemma homeomorphism_injective_open_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 429 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 430 | and imf: "f ` S = T" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 431 | and injf: "inj_on f S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 432 | 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 | 433 | obtains g where "homeomorphism S T f g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 434 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 435 | have "continuous_on T (inv_into S f)" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 436 | 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 | 437 | 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 | 438 | by (auto simp: homeomorphism_def) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 439 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 440 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 441 | lemma homeomorphism_injective_closed_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 442 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 443 | and imf: "f ` S = T" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 444 | and injf: "inj_on f S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 445 | 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 | 446 | obtains g where "homeomorphism S T f g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 447 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 448 | have "continuous_on T (inv_into S f)" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 449 | 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 | 450 | 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 | 451 | by (auto simp: homeomorphism_def) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 452 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 453 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 454 | lemma homeomorphism_imp_open_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 455 | 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: 
69753diff
changeset | 456 | 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: 
69753diff
changeset | 457 | shows "openin (top_of_set T) (f ` U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 458 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 459 | 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 | 460 | 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 | 461 | from hom have "continuous_on T g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 462 | unfolding homeomorphism_def by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 463 | moreover have "g ` T = S" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 464 | by (metis hom homeomorphism_def) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 465 | ultimately show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 466 | by (simp add: continuous_on_open oo) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 467 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 468 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 469 | lemma homeomorphism_imp_closed_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 470 | 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: 
69753diff
changeset | 471 | 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: 
69753diff
changeset | 472 | shows "closedin (top_of_set T) (f ` U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 473 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 474 | 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 | 475 | 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 | 476 | from hom have "continuous_on T g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 477 | unfolding homeomorphism_def by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 478 | moreover have "g ` T = S" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 479 | by (metis hom homeomorphism_def) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 480 | ultimately show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 481 | by (simp add: continuous_on_closed oo) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 482 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 483 | |
| 70136 | 484 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Seperability\<close> | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 485 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 486 | lemma subset_second_countable: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 487 | obtains \<B> :: "'a:: second_countable_topology set set" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 488 | where "countable \<B>" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 489 |           "{} \<notin> \<B>"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 490 | "\<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: 
69753diff
changeset | 491 | "\<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 \<B> :: "'a set set" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 494 | where "countable \<B>" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 495 | 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: 
69753diff
changeset | 496 | 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 | 497 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 498 | obtain \<C> :: "'a set set" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 499 | 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 | 500 | 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 | 501 | by (metis univ_second_countable that) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 502 | show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 503 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 504 | show "countable ((\<lambda>C. S \<inter> C) ` \<C>)" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 505 | 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: 
69753diff
changeset | 506 | 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 | 507 | using ope by auto | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 508 | 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 | 509 | by (metis \<C> image_mono inf_Sup openin_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 510 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 511 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 512 | show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 513 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 514 |     show "countable (\<B> - {{}})"
 | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 515 | 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: 
69753diff
changeset | 516 |     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: 
69753diff
changeset | 517 | 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: 
69753diff
changeset | 518 |     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 | 519 | using \<B> [OF that] | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 520 | apply clarify | 
| 76894 | 521 |       by (rule_tac x="\<U> - {{}}" in exI, auto)
 | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 522 | qed auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 523 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 524 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 525 | lemma Lindelof_openin: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 526 | 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: 
69753diff
changeset | 527 | 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 | 528 | 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 | 529 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 530 | 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 | 531 | using assms by (simp add: openin_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 532 | 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 | 533 | by metis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 534 | 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 | 535 | using tf by fastforce | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 536 | 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 | 537 | using tf by (force intro: Lindelof [of "tf ` \<F>"]) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 538 | 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 | 539 | by (clarsimp simp add: countable_subset_image) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 540 | then show ?thesis .. | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 541 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 542 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 543 | |
| 70136 | 544 | 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 | 545 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 546 | lemma continuous_imp_closed_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 547 | 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: 
69753diff
changeset | 548 | assumes "closedin (top_of_set S) U" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 549 | "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: 
69753diff
changeset | 550 | shows "closedin (top_of_set T) (f ` U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 551 | 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 | 552 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 553 | lemma closed_map_restrict: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 554 | 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: 
69753diff
changeset | 555 | 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 | 556 | and "T' \<subseteq> T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 557 | shows "closedin (top_of_set T') (f ` U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 558 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 559 | 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 | 560 | using cloU by (auto simp: closedin_closed) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 561 | 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 | 562 | by (fastforce simp add: closedin_closed) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 563 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 564 | |
| 70136 | 565 | 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 | 566 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 567 | lemma open_map_restrict: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 568 | 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: 
69753diff
changeset | 569 | 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 | 570 | and "T' \<subseteq> T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 571 | shows "openin (top_of_set T') (f ` U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 572 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 573 | 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 | 574 | using opeU by (auto simp: openin_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 575 | 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 | 576 | by (fastforce simp add: openin_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 577 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 578 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 579 | |
| 70136 | 580 | 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 | 581 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 582 | 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: 
78200diff
changeset | 583 | assumes T: "f \<in> S \<rightarrow> T" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 584 | 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: 
69753diff
changeset | 585 | \<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: 
69753diff
changeset | 586 | openin (top_of_set T) U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 587 | shows "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 588 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 589 | have [simp]: "S \<inter> f -` f ` S = S" by auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 590 | show ?thesis | 
| 72225 | 591 | 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 | 592 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 593 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 594 | 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: 
78200diff
changeset | 595 | assumes T: "f \<in> S \<rightarrow> T" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 596 | 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: 
69753diff
changeset | 597 | \<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: 
69753diff
changeset | 598 | closedin (top_of_set T) U)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 599 | shows "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 600 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 601 | have [simp]: "S \<inter> f -` f ` S = S" by auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 602 | show ?thesis | 
| 72225 | 603 | 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 | 604 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 605 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 606 | lemma open_map_imp_quotient_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 607 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 608 | and T: "T \<subseteq> f ` S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 609 | 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: 
69753diff
changeset | 610 | \<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: 
69753diff
changeset | 611 | 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: 
69753diff
changeset | 612 | openin (top_of_set (f ` S)) T" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 613 | proof - | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 614 | have "T = f ` (S \<inter> f -` T)" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 615 | using T by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 616 | then show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 617 | using "ope" contf continuous_on_open by metis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 618 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 619 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 620 | lemma closed_map_imp_quotient_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 621 | assumes contf: "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 622 | and T: "T \<subseteq> f ` S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 623 | 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: 
69753diff
changeset | 624 | \<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: 
69753diff
changeset | 625 | 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: 
69753diff
changeset | 626 | openin (top_of_set (f ` S)) T" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 627 | (is "?lhs = ?rhs") | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 628 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 629 | assume ?lhs | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 630 | 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 | 631 | using closedin_diff by fastforce | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 632 | 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 | 633 | using T by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 634 | show ?rhs | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 635 | using ope [OF *, unfolded closedin_def] by auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 636 | next | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 637 | assume ?rhs | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 638 | with contf show ?lhs | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 639 | by (auto simp: continuous_on_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 640 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 641 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 642 | 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: 
78200diff
changeset | 643 | 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: 
78200diff
changeset | 644 | 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 | 645 | 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 | 646 | and U: "U \<subseteq> T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 647 | 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: 
69753diff
changeset | 648 | openin (top_of_set T) U" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 649 | (is "?lhs = ?rhs") | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 650 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 651 | 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: 
69753diff
changeset | 652 | 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: 
69753diff
changeset | 653 | 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: 
69753diff
changeset | 654 | openin (top_of_set T) (T \<inter> g -` Z)" | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 655 | using contf contg by (auto simp: continuous_on_open) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 656 | show ?thesis | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 657 | proof | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 658 |     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 | 659 | using imf img by blast | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 660 | also have "... = U" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 661 | using U by auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 662 | 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 | 663 | assume ?lhs | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 664 | 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: 
78200diff
changeset | 665 | 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 | 666 | show ?rhs | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 667 | using g [OF *] eq by auto | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 668 | next | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 669 | assume rhs: ?rhs | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 670 | show ?lhs | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 671 | using assms continuous_openin_preimage rhs by blast | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 672 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 673 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 674 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 675 | lemma continuous_left_inverse_imp_quotient_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 676 | assumes "continuous_on S f" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 677 | and "continuous_on (f ` S) g" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 678 | 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 | 679 | and "U \<subseteq> f ` S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 680 | 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: 
69753diff
changeset | 681 | openin (top_of_set (f ` S)) U" | 
| 76894 | 682 | using assms | 
| 683 | by (intro continuous_right_inverse_imp_quotient_map) auto | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 684 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 685 | lemma continuous_imp_quotient_map: | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 686 | fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space" | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 687 | 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: 
69753diff
changeset | 688 | 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: 
69753diff
changeset | 689 | openin (top_of_set T) U" | 
| 76894 | 690 | 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 | 691 | |
| 70136 | 692 | 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 | 693 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 694 | subsubsection\<open>on open sets\<close> | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 695 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 696 | lemma pasting_lemma: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 697 | 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: 
69753diff
changeset | 698 | 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: 
69753diff
changeset | 699 | 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: 
69753diff
changeset | 700 | 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: 
69753diff
changeset | 701 | shows "continuous_map X Y g" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 702 | unfolding continuous_map_openin_preimage_eq | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 703 | 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: 
78200diff
changeset | 704 | show "g \<in> topspace X \<rightarrow> topspace Y" | 
| 71174 | 705 | 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: 
69753diff
changeset | 706 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 707 | fix U | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 708 | assume Y: "openin Y U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 709 | 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: 
69753diff
changeset | 710 | 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: 
69753diff
changeset | 711 | 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: 
69753diff
changeset | 712 | using f g T by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 713 | 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: 
69753diff
changeset | 714 | 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: 
69753diff
changeset | 715 | 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: 
69753diff
changeset | 716 | 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: 
69753diff
changeset | 717 | by (auto simp: *) | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 718 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 719 | |
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 720 | lemma pasting_lemma_exists: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 721 | 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: 
69753diff
changeset | 722 | 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: 
69753diff
changeset | 723 | 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: 
69753diff
changeset | 724 | 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: 
69753diff
changeset | 725 | 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 | 726 | proof | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 727 | 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: 
69753diff
changeset | 728 | show "continuous_map X Y ?h" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 729 | apply (rule pasting_lemma [OF ope cont]) | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 730 | apply (blast intro: f)+ | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 731 | 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: 
69753diff
changeset | 732 | 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: 
69753diff
changeset | 733 | 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: 
69753diff
changeset | 734 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 735 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 736 | lemma pasting_lemma_locally_finite: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 737 |   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: 
69753diff
changeset | 738 | 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: 
69753diff
changeset | 739 | 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: 
69753diff
changeset | 740 | 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: 
69753diff
changeset | 741 | 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: 
69753diff
changeset | 742 | shows "continuous_map X Y g" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 743 | unfolding continuous_map_closedin_preimage_eq | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 744 | 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: 
78200diff
changeset | 745 | show "g \<in> topspace X \<rightarrow> topspace Y" | 
| 71174 | 746 | 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: 
69753diff
changeset | 747 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 748 | fix U | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 749 | assume Y: "closedin Y U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 750 | 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: 
69753diff
changeset | 751 | 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: 
69753diff
changeset | 752 | 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: 
69753diff
changeset | 753 | using f g T by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 754 | 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: 
69753diff
changeset | 755 | 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: 
69753diff
changeset | 756 | 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: 
69753diff
changeset | 757 |   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: 
69753diff
changeset | 758 |            \<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: 
69753diff
changeset | 759 | by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 760 | 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: 
69753diff
changeset | 761 | using T by blast | 
| 76894 | 762 | 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: 
69753diff
changeset | 763 | unfolding locally_finite_in_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 764 | using finite_subset [OF sub] fin by force | 
| 76894 | 765 | then show "closedin X (topspace X \<inter> g -` U)" | 
| 766 | 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: 
69753diff
changeset | 767 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 768 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 769 | 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: 
69753diff
changeset | 770 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 771 | lemma pasting_lemma_closed: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 772 | assumes fin: "finite I" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 773 | 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: 
69753diff
changeset | 774 | 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: 
69753diff
changeset | 775 | 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: 
69753diff
changeset | 776 | 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: 
69753diff
changeset | 777 | shows "continuous_map X Y g" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 778 | 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: 
69753diff
changeset | 779 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 780 | lemma pasting_lemma_exists_locally_finite: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 781 |   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: 
69753diff
changeset | 782 | 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: 
69753diff
changeset | 783 | 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: 
69753diff
changeset | 784 | 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: 
69753diff
changeset | 785 | 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: 
69753diff
changeset | 786 | 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: 
69753diff
changeset | 787 | 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: 
69753diff
changeset | 788 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 789 | 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: 
69753diff
changeset | 790 | 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: 
69753diff
changeset | 791 | apply (blast intro: assms)+ | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 792 | 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 | 793 | next | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 794 | fix x i | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 795 | assume "i \<in> I" and "x \<in> topspace X \<inter> T i" | 
| 76894 | 796 | then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x" | 
| 797 | 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: 
69753diff
changeset | 798 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 799 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 800 | lemma pasting_lemma_exists_closed: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 801 | assumes fin: "finite I" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 802 | 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: 
69753diff
changeset | 803 | 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: 
69753diff
changeset | 804 | 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: 
69753diff
changeset | 805 | 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: 
69753diff
changeset | 806 | 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: 
69753diff
changeset | 807 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 808 | 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: 
69753diff
changeset | 809 | 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: 
69753diff
changeset | 810 | apply (blast intro: f)+ | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 811 | 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: 
69753diff
changeset | 812 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 813 | fix x i | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 814 | 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 | 815 | 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 | 816 | by (metis (no_types, lifting) IntD2 IntI f someI_ex) | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 817 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 818 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 819 | lemma continuous_map_cases: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 820 |   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: 
69753diff
changeset | 821 |       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: 
69753diff
changeset | 822 |       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: 
69753diff
changeset | 823 | 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: 
69753diff
changeset | 824 | proof (rule pasting_lemma_closed) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 825 | 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: 
69753diff
changeset | 826 | 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: 
69753diff
changeset | 827 |   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: 
69753diff
changeset | 828 |   show "finite {True,False}" by auto
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 829 |   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: 
69753diff
changeset | 830 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 831 | show "?f i x = ?f j x" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 832 |     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: 
69753diff
changeset | 833 | proof - | 
| 76894 | 834 | have "f x = g x" if "i" "\<not> j" | 
| 835 | by (smt (verit, best) Diff_Diff_Int closure_of_interior_of closure_of_restrict eq fg | |
| 836 | 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: 
69753diff
changeset | 837 | moreover | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 838 | have "g x = f x" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 839 |       if "x \<in> X closure_of {x. \<not> P x}" "x \<in> X closure_of Collect P" "\<not> i" "j" for x
 | 
| 76894 | 840 | 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: 
69753diff
changeset | 841 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 842 | 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: 
69753diff
changeset | 843 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 844 |   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: 
69753diff
changeset | 845 | if "x \<in> topspace X" for x | 
| 76894 | 846 | 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: 
69753diff
changeset | 847 | qed (auto simp: f g) | 
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 848 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 849 | lemma continuous_map_cases_alt: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 850 |   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: 
69753diff
changeset | 851 |       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: 
69753diff
changeset | 852 |       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: 
69753diff
changeset | 853 | 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: 
69753diff
changeset | 854 | apply (rule continuous_map_cases) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 855 | using assms | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 856 | 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: 
69753diff
changeset | 857 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 858 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 859 | lemma continuous_map_cases_function: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 860 | assumes contp: "continuous_map X Z p" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 861 |     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: 
69753diff
changeset | 862 |     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: 
69753diff
changeset | 863 | 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: 
69753diff
changeset | 864 | 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: 
69753diff
changeset | 865 | proof (rule continuous_map_cases_alt) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 866 |   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: 
69753diff
changeset | 867 | proof (rule continuous_map_from_subtopology_mono) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 868 |     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: 
69753diff
changeset | 869 | 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: 
69753diff
changeset | 870 | by (simp add: contf) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 871 |     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: 
69753diff
changeset | 872 | 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: 
69753diff
changeset | 873 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 874 |   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: 
69753diff
changeset | 875 | proof (rule continuous_map_from_subtopology_mono) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 876 |     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: 
69753diff
changeset | 877 | 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: 
69753diff
changeset | 878 | by (simp add: contg) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 879 |     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: 
78256diff
changeset | 880 | 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: 
69753diff
changeset | 881 |     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: 
69753diff
changeset | 882 | 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: 
69753diff
changeset | 883 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 884 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 885 |   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: 
69753diff
changeset | 886 | 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 | 887 | qed | 
| 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 888 | |
| 69750 | 889 | subsection \<open>Retractions\<close> | 
| 890 | ||
| 70136 | 891 | definition\<^marker>\<open>tag important\<close> retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 69750 | 892 | 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: 
78200diff
changeset | 893 | T \<subseteq> S \<and> continuous_on S r \<and> r \<in> S \<rightarrow> T \<and> (\<forall>x\<in>T. r x = x)" | 
| 69750 | 894 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78685diff
changeset | 895 | definition\<^marker>\<open>tag important\<close> retract_of (infixl \<open>retract'_of\<close> 50) where | 
| 69750 | 896 | "T retract_of S \<longleftrightarrow> (\<exists>r. retraction S T r)" | 
| 897 | ||
| 898 | lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow> r (r x) = r x" | |
| 899 | unfolding retraction_def by auto | |
| 900 | ||
| 901 | text \<open>Preservation of fixpoints under (more general notion of) retraction\<close> | |
| 902 | ||
| 903 | lemma invertible_fixpoint_property: | |
| 904 | fixes S :: "'a::topological_space set" | |
| 905 | and T :: "'b::topological_space set" | |
| 906 | 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: 
78200diff
changeset | 907 | and "i \<in> T \<rightarrow> S" | 
| 69750 | 908 | 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: 
78200diff
changeset | 909 | and "r \<in> S \<rightarrow> T" | 
| 69750 | 910 | 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: 
78200diff
changeset | 911 | and FP: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" | 
| 69750 | 912 | 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: 
78200diff
changeset | 913 | and "g \<in> T \<rightarrow> T" | 
| 69750 | 914 | obtains y where "y \<in> T" and "g y = y" | 
| 915 | proof - | |
| 916 | have "\<exists>x\<in>S. (i \<circ> g \<circ> r) x = x" | |
| 917 | proof (rule FP) | |
| 918 | 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: 
78200diff
changeset | 919 | 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: 
78200diff
changeset | 920 | show "(i \<circ> g \<circ> r) \<in> S \<rightarrow> S" | 
| 69750 | 921 | using assms(2,4,8) by force | 
| 922 | qed | |
| 923 | then obtain x where x: "x \<in> S" "(i \<circ> g \<circ> r) x = x" .. | |
| 924 | then have *: "g (r x) \<in> T" | |
| 925 | using assms(4,8) by auto | |
| 926 | have "r ((i \<circ> g \<circ> r) x) = r x" | |
| 927 | using x by auto | |
| 928 | then show ?thesis | |
| 929 | using "*" ri that by auto | |
| 930 | qed | |
| 931 | ||
| 932 | lemma homeomorphic_fixpoint_property: | |
| 933 | fixes S :: "'a::topological_space set" | |
| 934 | and T :: "'b::topological_space set" | |
| 935 | assumes "S homeomorphic T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 936 | 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: 
78200diff
changeset | 937 | (\<forall>g. continuous_on T g \<and> g \<in> T \<rightarrow> T \<longrightarrow> (\<exists>y\<in>T. g y = y))" | 
| 69750 | 938 | (is "?lhs = ?rhs") | 
| 939 | proof - | |
| 940 | obtain r i where r: | |
| 941 | "\<forall>x\<in>S. i (r x) = x" "r ` S = T" "continuous_on S r" | |
| 942 | "\<forall>y\<in>T. r (i y) = y" "i ` T = S" "continuous_on T i" | |
| 943 | using assms unfolding homeomorphic_def homeomorphism_def by blast | |
| 944 | show ?thesis | |
| 945 | proof | |
| 946 | assume ?lhs | |
| 947 | with r show ?rhs | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 948 | by (smt (verit, del_insts) Pi_iff image_eqI invertible_fixpoint_property[of T i S r]) | 
| 69750 | 949 | next | 
| 950 | assume ?rhs | |
| 951 | with r show ?lhs | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 952 | by (smt (verit, del_insts) Pi_iff image_eqI invertible_fixpoint_property[of S r T i]) | 
| 69750 | 953 | qed | 
| 954 | qed | |
| 955 | ||
| 956 | lemma retract_fixpoint_property: | |
| 957 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | |
| 958 | and S :: "'a set" | |
| 959 | 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: 
78200diff
changeset | 960 | and FP: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" | 
| 69750 | 961 | 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: 
78200diff
changeset | 962 | and "g \<in> T \<rightarrow> T" | 
| 69750 | 963 | obtains y where "y \<in> T" and "g y = y" | 
| 964 | proof - | |
| 965 | obtain h where "retraction S T h" | |
| 966 | using assms(1) unfolding retract_of_def .. | |
| 967 | then show ?thesis | |
| 968 | unfolding retraction_def | |
| 969 | 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: 
78200diff
changeset | 970 | by (smt (verit, del_insts) Pi_iff assms(4) contg subsetD that) | 
| 69750 | 971 | qed | 
| 972 | ||
| 973 | lemma retraction: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 974 | "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: 
78200diff
changeset | 975 | by (force simp: retraction_def simp flip: image_subset_iff_funcset) | 
| 69750 | 976 | |
| 69753 | 977 | lemma retractionE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close> | 
| 69750 | 978 | 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: 
78200diff
changeset | 979 | 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 | 980 | proof (rule that) | 
| 981 | from retraction [of S T r] assms | |
| 982 | have "T \<subseteq> S" "continuous_on S r" "r ` S = T" and "\<forall>x \<in> T. r x = x" | |
| 983 | by simp_all | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 984 | 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: 
78200diff
changeset | 985 | by auto | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 986 | then show "T = r ` S" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 987 | using \<open>r ` S = T\<close> by blast | 
| 69750 | 988 | from \<open>\<forall>x \<in> T. r x = x\<close> have "r x = x" if "x \<in> T" for x | 
| 989 | using that by simp | |
| 990 | with \<open>r ` S = T\<close> show "r (r x) = r x" if "x \<in> S" for x | |
| 991 | using that by auto | |
| 992 | qed | |
| 993 | ||
| 69753 | 994 | lemma retract_ofE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close> | 
| 69750 | 995 | 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: 
78200diff
changeset | 996 | 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 | 997 | proof - | 
| 998 | from assms obtain r where "retraction S T r" | |
| 999 | by (auto simp add: retract_of_def) | |
| 1000 | with that show thesis | |
| 1001 | by (auto elim: retractionE) | |
| 1002 | qed | |
| 1003 | ||
| 1004 | 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: 
78200diff
changeset | 1005 | 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: 
78200diff
changeset | 1006 | obtains g where "continuous_on T g" "g \<in> T \<rightarrow> U" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" | 
| 69750 | 1007 | proof - | 
| 1008 | from \<open>S retract_of T\<close> obtain r where "retraction T S r" | |
| 1009 | 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: 
78200diff
changeset | 1010 | 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: 
78200diff
changeset | 1011 | 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: 
78200diff
changeset | 1012 | then show thesis | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1013 | by (smt (verit, best) Pi_iff \<open>retraction T S r\<close> assms(3) comp_apply retraction_def that) | 
| 69750 | 1014 | qed | 
| 1015 | ||
| 1016 | lemma idempotent_imp_retraction: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1017 | assumes "continuous_on S f" and "f \<in> S \<rightarrow> S" and "\<And>x. x \<in> S \<Longrightarrow> f(f x) = f x" | 
| 69750 | 1018 | 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: 
78200diff
changeset | 1019 | by (simp add: assms funcset_image retraction) | 
| 69750 | 1020 | |
| 1021 | lemma retraction_subset: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1022 | 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: 
78200diff
changeset | 1023 | shows "retraction S' T r" | 
| 69750 | 1024 | unfolding retraction_def | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1025 | by (metis assms continuous_on_subset image_mono image_subset_iff_funcset retraction) | 
| 69750 | 1026 | |
| 1027 | lemma retract_of_subset: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1028 | 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: 
78200diff
changeset | 1029 | shows "T retract_of S'" | 
| 69750 | 1030 | by (meson assms retract_of_def retraction_subset) | 
| 1031 | ||
| 1032 | lemma retraction_refl [simp]: "retraction S S (\<lambda>x. x)" | |
| 1033 | by (simp add: retraction) | |
| 1034 | ||
| 1035 | lemma retract_of_refl [iff]: "S retract_of S" | |
| 1036 | unfolding retract_of_def retraction_def | |
| 1037 | using continuous_on_id by blast | |
| 1038 | ||
| 1039 | lemma retract_of_imp_subset: | |
| 1040 | "S retract_of T \<Longrightarrow> S \<subseteq> T" | |
| 1041 | by (simp add: retract_of_def retraction_def) | |
| 1042 | ||
| 1043 | lemma retract_of_empty [simp]: | |
| 1044 |      "({} retract_of S) \<longleftrightarrow> S = {}"  "(S retract_of {}) \<longleftrightarrow> S = {}"
 | |
| 1045 | by (auto simp: retract_of_def retraction_def) | |
| 1046 | ||
| 1047 | lemma retract_of_singleton [iff]: "({x} retract_of S) \<longleftrightarrow> x \<in> S"
 | |
| 1048 | unfolding retract_of_def retraction_def by force | |
| 1049 | ||
| 1050 | lemma retraction_comp: | |
| 76894 | 1051 | "\<lbrakk>retraction S T f; retraction T U g\<rbrakk> \<Longrightarrow> retraction S U (g \<circ> f)" | 
| 1052 | by (smt (verit, best) comp_apply continuous_on_compose image_comp retraction subset_iff) | |
| 69750 | 1053 | |
| 1054 | lemma retract_of_trans [trans]: | |
| 1055 | assumes "S retract_of T" and "T retract_of U" | |
| 1056 | shows "S retract_of U" | |
| 1057 | using assms by (auto simp: retract_of_def intro: retraction_comp) | |
| 1058 | ||
| 1059 | lemma closedin_retract: | |
| 1060 | fixes S :: "'a :: t2_space set" | |
| 1061 | assumes "S retract_of T" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1062 | shows "closedin (top_of_set T) S" | 
| 69750 | 1063 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1064 | 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 | 1065 | using assms by (auto simp: retract_of_def retraction_def) | 
| 1066 |   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: 
78200diff
changeset | 1067 | using r by force | 
| 69750 | 1068 |   also have "\<dots> = T \<inter> ((\<lambda>x. (x, r x)) -` ({y. \<exists>x. y = (x, x)}))"
 | 
| 1069 | unfolding vimage_def mem_Times_iff fst_conv snd_conv | |
| 1070 | using r | |
| 1071 | by auto | |
| 1072 | also have "closedin (top_of_set T) \<dots>" | |
| 1073 | by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) | |
| 1074 | finally show ?thesis . | |
| 1075 | qed | |
| 1076 | ||
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1077 | lemma closedin_self [simp]: "closedin (top_of_set S) S" | 
| 69750 | 1078 | by simp | 
| 1079 | ||
| 1080 | lemma retract_of_closed: | |
| 1081 | fixes S :: "'a :: t2_space set" | |
| 1082 | shows "\<lbrakk>closed T; S retract_of T\<rbrakk> \<Longrightarrow> closed S" | |
| 1083 | by (metis closedin_retract closedin_closed_eq) | |
| 1084 | ||
| 1085 | lemma retract_of_compact: | |
| 1086 | "\<lbrakk>compact T; S retract_of T\<rbrakk> \<Longrightarrow> compact S" | |
| 1087 | by (metis compact_continuous_image retract_of_def retraction) | |
| 1088 | ||
| 1089 | lemma retract_of_connected: | |
| 1090 | "\<lbrakk>connected T; S retract_of T\<rbrakk> \<Longrightarrow> connected S" | |
| 1091 | by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) | |
| 1092 | ||
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1093 | lemma retraction_openin_vimage_iff: | 
| 76894 | 1094 | assumes r: "retraction S T r" and "U \<subseteq> T" | 
| 1095 | shows "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U" (is "?lhs = ?rhs") | |
| 1096 | proof | |
| 1097 | show "?lhs \<Longrightarrow> ?rhs" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1098 | using r | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1099 | by (smt (verit, del_insts) assms(2) continuous_right_inverse_imp_quotient_map image_subset_iff_funcset r retractionE retraction_def retraction_subset) | 
| 76894 | 1100 | show "?rhs \<Longrightarrow> ?lhs" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1101 | by (metis continuous_on_open r retraction) | 
| 76894 | 1102 | qed | 
| 69750 | 1103 | |
| 1104 | lemma retract_of_Times: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1105 | "\<lbrakk>S retract_of S'; T retract_of T'\<rbrakk> \<Longrightarrow> (S \<times> T) retract_of (S' \<times> T')" | 
| 69750 | 1106 | apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) | 
| 1107 | apply (rename_tac f g) | |
| 1108 | apply (rule_tac x="\<lambda>z. ((f \<circ> fst) z, (g \<circ> snd) z)" in exI) | |
| 1109 | apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ | |
| 1110 | done | |
| 1111 | ||
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1112 | 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: 
69753diff
changeset | 1113 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78685diff
changeset | 1114 | definition retract_of_space :: "'a set \<Rightarrow> 'a topology \<Rightarrow> bool" (infix \<open>retract'_of'_space\<close> 50) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1115 | where "S retract_of_space X | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1116 | \<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: 
69753diff
changeset | 1117 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1118 | lemma retract_of_space_retraction_maps: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1119 | "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: 
69753diff
changeset | 1120 | 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: 
69753diff
changeset | 1121 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1122 | lemma retract_of_space_section_map: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1123 | "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: 
69753diff
changeset | 1124 | 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: 
69753diff
changeset | 1125 | by (auto simp: continuous_map_from_subtopology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1126 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1127 | lemma retract_of_space_imp_subset: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1128 | "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: 
69753diff
changeset | 1129 | by (simp add: retract_of_space_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1130 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1131 | lemma retract_of_space_topspace: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1132 | "topspace X retract_of_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1133 | using retract_of_space_def by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1134 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1135 | lemma retract_of_space_empty [simp]: | 
| 78336 | 1136 |    "{} retract_of_space X \<longleftrightarrow> X = trivial_topology"
 | 
| 1137 | by (auto simp: retract_of_space_def) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1138 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1139 | lemma retract_of_space_singleton [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1140 |   "{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: 
69753diff
changeset | 1141 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1142 |   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: 
69753diff
changeset | 1143 | using that by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1144 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1145 | by (force simp: retract_of_space_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1146 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1147 | |
| 77934 | 1148 | lemma retract_of_space_trans: | 
| 1149 | assumes "S retract_of_space X" "T retract_of_space subtopology X S" | |
| 1150 | shows "T retract_of_space X" | |
| 1151 | using assms | |
| 1152 | apply (simp add: retract_of_space_retraction_maps) | |
| 1153 | by (metis id_comp inf.absorb_iff2 retraction_maps_compose subtopology_subtopology) | |
| 1154 | ||
| 1155 | lemma retract_of_space_subtopology: | |
| 1156 | assumes "S retract_of_space X" "S \<subseteq> U" | |
| 1157 | shows "S retract_of_space subtopology X U" | |
| 1158 | using assms | |
| 1159 | apply (clarsimp simp: retract_of_space_def) | |
| 1160 | by (metis continuous_map_from_subtopology inf.absorb2 subtopology_subtopology) | |
| 1161 | ||
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1162 | lemma retract_of_space_clopen: | 
| 78336 | 1163 |   assumes "openin X S" "closedin X S" "S = {} \<Longrightarrow> X = trivial_topology"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1164 | shows "S retract_of_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1165 | proof (cases "S = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1166 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1167 | then obtain a where "a \<in> S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1168 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1169 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1170 | unfolding retract_of_space_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1171 | proof (intro exI conjI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1172 | show "S \<subseteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1173 | by (simp add: assms closedin_subset) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1174 | 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: 
69753diff
changeset | 1175 | proof (rule continuous_map_cases) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1176 |       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: 
69753diff
changeset | 1177 | by (simp add: continuous_map_from_subtopology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1178 |       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: 
69753diff
changeset | 1179 | 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: 
69753diff
changeset | 1180 |       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: 
69753diff
changeset | 1181 | 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: 
69753diff
changeset | 1182 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1183 | 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: 
69753diff
changeset | 1184 | 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: 
69753diff
changeset | 1185 | qed auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1186 | qed (use assms in auto) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1187 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1188 | lemma retract_of_space_disjoint_union: | 
| 78336 | 1189 |   assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \<union> T = topspace X" and "S = {} \<Longrightarrow> X = trivial_topology"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1190 | shows "S retract_of_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1191 | proof (rule retract_of_space_clopen) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1192 |   have "S \<inter> T = {}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1193 | by (meson ST disjnt_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1194 | then have "S = topspace X - T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1195 | using ST by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1196 | then show "closedin X S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1197 | 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: 
69753diff
changeset | 1198 | qed (auto simp: assms) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1199 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1200 | lemma retraction_maps_section_image1: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1201 | assumes "retraction_maps X Y r s" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1202 | 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: 
69753diff
changeset | 1203 | unfolding retract_of_space_section_map | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1204 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1205 | show "s ` topspace Y \<subseteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1206 | 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: 
69753diff
changeset | 1207 | 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: 
69753diff
changeset | 1208 | unfolding section_map_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1209 | 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: 
69753diff
changeset | 1210 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1211 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1212 | lemma retraction_maps_section_image2: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1213 | "retraction_maps X Y r s | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1214 | \<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: 
69753diff
changeset | 1215 | 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: 
69753diff
changeset | 1216 | section_map_def by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1217 | |
| 77934 | 1218 | lemma hereditary_imp_retractive_property: | 
| 1219 | assumes "\<And>X S. P X \<Longrightarrow> P(subtopology X S)" | |
| 1220 | "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')" | |
| 1221 | assumes "retraction_map X X' r" "P X" | |
| 1222 | shows "Q X'" | |
| 1223 | by (meson assms retraction_map_def retraction_maps_section_image2) | |
| 1224 | ||
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1225 | subsection\<open>Paths and path-connectedness\<close> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1226 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1227 | 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: 
69753diff
changeset | 1228 |    "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: 
69753diff
changeset | 1229 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1230 | lemma pathin_compose: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1231 | "\<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: 
69753diff
changeset | 1232 | 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: 
69753diff
changeset | 1233 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1234 | lemma pathin_subtopology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1235 |      "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: 
69753diff
changeset | 1236 | 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: 
69753diff
changeset | 1237 | |
| 78336 | 1238 | lemma pathin_const [simp]: "pathin X (\<lambda>x. a) \<longleftrightarrow> a \<in> topspace X" | 
| 1239 | using pathin_subtopology by (fastforce simp add: pathin_def) | |
| 1240 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1241 | 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: 
69922diff
changeset | 1242 | 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: 
69922diff
changeset | 1243 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1244 | 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: 
69922diff
changeset | 1245 | 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: 
69922diff
changeset | 1246 | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1247 | 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: 
69922diff
changeset | 1248 | 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: 
69753diff
changeset | 1249 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1250 | 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: 
69753diff
changeset | 1251 | 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: 
69753diff
changeset | 1252 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1253 | 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: 
69753diff
changeset | 1254 | 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: 
69753diff
changeset | 1255 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1256 | lemma path_connectedin_absolute [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1257 | "path_connectedin (subtopology X S) S \<longleftrightarrow> path_connectedin X S" | 
| 71172 | 1258 | 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: 
69753diff
changeset | 1259 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1260 | lemma path_connectedin_subset_topspace: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1261 | "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: 
69753diff
changeset | 1262 | by (simp add: path_connectedin_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1263 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1264 | lemma path_connectedin_subtopology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1265 | "path_connectedin (subtopology X S) T \<longleftrightarrow> path_connectedin X T \<and> T \<subseteq> S" | 
| 71172 | 1266 | 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: 
69753diff
changeset | 1267 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1268 | lemma path_connectedin: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1269 | "path_connectedin X S \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1270 | 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: 
78200diff
changeset | 1271 |         (\<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: 
69753diff
changeset | 1272 | 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: 
78200diff
changeset | 1273 | 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: 
69753diff
changeset | 1274 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1275 | lemma path_connectedin_topspace: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1276 | "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: 
69753diff
changeset | 1277 | by (simp add: path_connectedin_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1278 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1279 | lemma path_connected_imp_connected_space: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1280 | assumes "path_connected_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1281 | shows "connected_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1282 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1283 | 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: 
69753diff
changeset | 1284 | proof (intro exI conjI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1285 |     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: 
69753diff
changeset | 1286 | 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: 
69753diff
changeset | 1287 |     then show "connectedin X (g ` {0..1})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1288 | by (rule connectedin_continuous_map_image) auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1289 | qed auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1290 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1291 | using assms | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1292 | 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: 
69753diff
changeset | 1293 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1294 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1295 | lemma path_connectedin_imp_connectedin: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1296 | "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: 
69753diff
changeset | 1297 | 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: 
69753diff
changeset | 1298 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1299 | lemma path_connected_space_topspace_empty: | 
| 78336 | 1300 | "path_connected_space trivial_topology" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1301 | by (simp add: path_connected_space_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1302 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1303 | 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: 
69753diff
changeset | 1304 | by (simp add: path_connectedin) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1305 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1306 | 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: 
69753diff
changeset | 1307 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1308 |   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: 
69753diff
changeset | 1309 | by (simp add: path_connectedin) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1310 |   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: 
69753diff
changeset | 1311 | unfolding path_connectedin | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1312 | using pathin_const by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1313 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1314 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1315 | lemma path_connectedin_continuous_map_image: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1316 | 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: 
69753diff
changeset | 1317 | shows "path_connectedin Y (f ` S)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1318 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1319 | 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: 
78200diff
changeset | 1320 | 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: 
69753diff
changeset | 1321 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1322 | unfolding path_connectedin | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1323 | proof (intro conjI ballI; clarify?) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1324 | fix x | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1325 | assume "x \<in> S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1326 | 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: 
78200diff
changeset | 1327 | 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: 
69753diff
changeset | 1328 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1329 | fix x y | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1330 | 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: 
78200diff
changeset | 1331 |     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: 
69753diff
changeset | 1332 | 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: 
78200diff
changeset | 1333 |     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: 
69753diff
changeset | 1334 | proof (intro exI conjI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1335 | show "pathin Y (f \<circ> g)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1336 | 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: 
69753diff
changeset | 1337 | qed (use g in auto) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1338 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1339 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1340 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1341 | lemma path_connectedin_discrete_topology: | 
| 76894 | 1342 |   "path_connectedin (discrete_topology U) S \<longleftrightarrow> S \<subseteq> U \<and> (\<exists>a. S \<subseteq> {a})" (is "?lhs = ?rhs")
 | 
| 1343 | proof | |
| 1344 | show "?lhs \<Longrightarrow> ?rhs" | |
| 1345 | by (meson connectedin_discrete_topology path_connectedin_imp_connectedin) | |
| 1346 | show "?rhs \<Longrightarrow> ?lhs" | |
| 1347 | using subset_singletonD by fastforce | |
| 1348 | qed | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1349 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1350 | 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: 
69922diff
changeset | 1351 |    "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: 
69922diff
changeset | 1352 | 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: 
69922diff
changeset | 1353 | 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: 
69922diff
changeset | 1354 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1355 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1356 | lemma homeomorphic_path_connected_space_imp: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1357 | "\<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: 
69753diff
changeset | 1358 | 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: 
78256diff
changeset | 1359 | 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: 
69753diff
changeset | 1360 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1361 | lemma homeomorphic_path_connected_space: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1362 | "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: 
69753diff
changeset | 1363 | 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: 
69753diff
changeset | 1364 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1365 | lemma homeomorphic_map_path_connectedness: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1366 | 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: 
69753diff
changeset | 1367 | 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: 
69753diff
changeset | 1368 | unfolding path_connectedin_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1369 | 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: 
78200diff
changeset | 1370 | 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: 
69753diff
changeset | 1371 | 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: 
69753diff
changeset | 1372 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1373 | assume "U \<subseteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1374 | 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: 
69753diff
changeset | 1375 | using assms unfolding homeomorphic_eq_everything_map | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72225diff
changeset | 1376 | 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: 
69753diff
changeset | 1377 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1378 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1379 | lemma homeomorphic_map_path_connectedness_eq: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1380 | "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: 
69753diff
changeset | 1381 | 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: 
69753diff
changeset | 1382 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1383 | subsection\<open>Connected components\<close> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1384 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1385 | 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: 
69753diff
changeset | 1386 | 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: 
69753diff
changeset | 1387 | \<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: 
69753diff
changeset | 1388 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1389 | abbreviation connected_component_of_set | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1390 | 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: 
69753diff
changeset | 1391 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1392 | 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: 
69753diff
changeset | 1393 | 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: 
69753diff
changeset | 1394 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1395 | lemma connected_component_in_topspace: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1396 | "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: 
69753diff
changeset | 1397 | 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: 
69753diff
changeset | 1398 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1399 | lemma connected_component_of_refl: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1400 | "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: 
69753diff
changeset | 1401 | 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: 
69753diff
changeset | 1402 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1403 | lemma connected_component_of_sym: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1404 | "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: 
69753diff
changeset | 1405 | by (meson connected_component_of_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1406 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1407 | lemma connected_component_of_trans: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1408 | "\<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: 
69753diff
changeset | 1409 | \<Longrightarrow> connected_component_of X x z" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1410 | unfolding connected_component_of_def | 
| 71842 
db120661dded
new HOL simproc: eliminate_false_implies
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1411 | using connectedin_Un by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1412 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1413 | lemma connected_component_of_mono: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1414 | "\<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: 
69753diff
changeset | 1415 | \<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: 
69753diff
changeset | 1416 | 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: 
69753diff
changeset | 1417 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1418 | lemma connected_component_of_set: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1419 |    "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: 
69753diff
changeset | 1420 | by (meson connected_component_of_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1421 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1422 | lemma connected_component_of_subset_topspace: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1423 | "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: 
69753diff
changeset | 1424 | using connected_component_in_topspace by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1425 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1426 | lemma connected_component_of_eq_empty: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1427 |    "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: 
69753diff
changeset | 1428 | 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: 
69753diff
changeset | 1429 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1430 | lemma connected_space_iff_connected_component: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1431 | "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: 
69753diff
changeset | 1432 | 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: 
69753diff
changeset | 1433 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1434 | lemma connected_space_imp_connected_component_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1435 | "\<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: 
69753diff
changeset | 1436 | \<Longrightarrow> connected_component_of X a b" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1437 | 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: 
69753diff
changeset | 1438 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1439 | lemma connected_space_connected_component_set: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1440 | "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: 
69753diff
changeset | 1441 | 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: 
69753diff
changeset | 1442 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1443 | lemma connected_component_of_maximal: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1444 | "\<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: 
69753diff
changeset | 1445 | 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: 
69753diff
changeset | 1446 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1447 | lemma connected_component_of_equiv: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1448 | "connected_component_of X x y \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1449 | 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: 
69753diff
changeset | 1450 | 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: 
69753diff
changeset | 1451 | 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: 
69753diff
changeset | 1452 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1453 | lemma connected_component_of_disjoint: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1454 | "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: 
69753diff
changeset | 1455 | \<longleftrightarrow> ~(connected_component_of X x y)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1456 | 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: 
69753diff
changeset | 1457 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1458 | lemma connected_component_of_eq: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1459 | "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: 
69753diff
changeset | 1460 | (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: 
69753diff
changeset | 1461 | 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: 
69753diff
changeset | 1462 | connected_component_of X x y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1463 | 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: 
69753diff
changeset | 1464 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1465 | lemma connectedin_connected_component_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1466 | "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: 
69753diff
changeset | 1467 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1468 |   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: 
69753diff
changeset | 1469 | by (auto simp: connected_component_of_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1470 | then show ?thesis | 
| 76894 | 1471 | 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: 
69753diff
changeset | 1472 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1473 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1474 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1475 | lemma Union_connected_components_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1476 | "\<Union>(connected_components_of X) = topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1477 | unfolding connected_components_of_def | 
| 76894 | 1478 | 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: 
69753diff
changeset | 1479 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1480 | lemma connected_components_of_maximal: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1481 | "\<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: 
69753diff
changeset | 1482 | unfolding connected_components_of_def disjnt_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1483 | apply clarify | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1484 | 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: 
69753diff
changeset | 1485 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1486 | lemma pairwise_disjoint_connected_components_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1487 | "pairwise disjnt (connected_components_of X)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1488 | unfolding connected_components_of_def pairwise_def | 
| 76894 | 1489 | 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: 
69753diff
changeset | 1490 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1491 | lemma complement_connected_components_of_Union: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1492 | "C \<in> connected_components_of X | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1493 |       \<Longrightarrow> topspace X - C = \<Union> (connected_components_of X - {C})"
 | 
| 76894 | 1494 | by (metis Union_connected_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint | 
| 1495 | 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: 
69753diff
changeset | 1496 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1497 | lemma nonempty_connected_components_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1498 |    "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: 
69753diff
changeset | 1499 | unfolding connected_components_of_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1500 | 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: 
69753diff
changeset | 1501 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1502 | lemma connected_components_of_subset: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1503 | "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: 
69753diff
changeset | 1504 | using Union_connected_components_of by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1505 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1506 | lemma connectedin_connected_components_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1507 | assumes "C \<in> connected_components_of X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1508 | shows "connectedin X C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1509 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1510 | 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: 
69753diff
changeset | 1511 | 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: 
69753diff
changeset | 1512 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1513 | using connectedin_connected_component_of by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1514 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1515 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1516 | lemma connected_component_in_connected_components_of: | 
| 76894 | 1517 | "connected_component_of_set X a \<in> connected_components_of X \<longleftrightarrow> a \<in> topspace X" | 
| 1518 | 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: 
69753diff
changeset | 1519 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1520 | lemma connected_space_iff_components_eq: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1521 | "connected_space X \<longleftrightarrow> (\<forall>C \<in> connected_components_of X. \<forall>C' \<in> connected_components_of X. C = C')" | 
| 76894 | 1522 | (is "?lhs = ?rhs") | 
| 1523 | proof | |
| 1524 | show "?lhs \<Longrightarrow> ?rhs" | |
| 1525 | by (simp add: connected_components_of_def connected_space_connected_component_set) | |
| 1526 | show "?rhs \<Longrightarrow> ?lhs" | |
| 1527 | by (metis Union_connected_components_of Union_iff connected_space_subconnected connectedin_connected_components_of) | |
| 1528 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1529 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1530 | lemma connected_components_of_eq_empty: | 
| 78336 | 1531 |    "connected_components_of X = {} \<longleftrightarrow> X = trivial_topology"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1532 | by (simp add: connected_components_of_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1533 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1534 | lemma connected_components_of_empty_space: | 
| 78336 | 1535 |    "connected_components_of trivial_topology = {}"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1536 | 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: 
69753diff
changeset | 1537 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1538 | lemma connected_components_of_subset_sing: | 
| 78336 | 1539 |    "connected_components_of X \<subseteq> {S} \<longleftrightarrow> connected_space X \<and> (X = trivial_topology \<or> topspace X = S)"
 | 
| 1540 | proof (cases "X = trivial_topology") | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1541 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1542 | then show ?thesis | 
| 78336 | 1543 | by (simp add: connected_components_of_empty_space) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1544 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1545 | case False | 
| 76894 | 1546 |   then have "\<lbrakk>connected_components_of X \<subseteq> {S}\<rbrakk> \<Longrightarrow> topspace X = S"
 | 
| 78336 | 1547 | using Union_connected_components_of connected_components_of_eq_empty by fastforce | 
| 76894 | 1548 | with False show ?thesis | 
| 1549 | unfolding connected_components_of_def | |
| 1550 | 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: 
69753diff
changeset | 1551 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1552 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1553 | lemma connected_space_iff_components_subset_singleton: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1554 |    "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: 
69753diff
changeset | 1555 | 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: 
69753diff
changeset | 1556 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1557 | lemma connected_components_of_eq_singleton: | 
| 78336 | 1558 |    "connected_components_of X = {S} \<longleftrightarrow> connected_space X \<and> X \<noteq> trivial_topology \<and> S = topspace X"
 | 
| 1559 | by (metis connected_components_of_eq_empty connected_components_of_subset_sing insert_not_empty subset_singleton_iff) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1560 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1561 | lemma connected_components_of_connected_space: | 
| 78336 | 1562 |    "connected_space X \<Longrightarrow> connected_components_of X = (if X = trivial_topology then {} else {topspace X})"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1563 | 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: 
69753diff
changeset | 1564 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1565 | lemma exists_connected_component_of_superset: | 
| 78336 | 1566 | assumes "connectedin X S" and ne: "X \<noteq> trivial_topology" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1567 | 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: 
69753diff
changeset | 1568 | proof (cases "S = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1569 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1570 | then show ?thesis | 
| 78336 | 1571 | using ne connected_components_of_eq_empty by fastforce | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1572 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1573 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1574 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1575 | 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: 
69753diff
changeset | 1576 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1577 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1578 | lemma closedin_connected_components_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1579 | assumes "C \<in> connected_components_of X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1580 | shows "closedin X C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1581 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1582 | 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: 
69753diff
changeset | 1583 | 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: 
69753diff
changeset | 1584 | 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: 
69753diff
changeset | 1585 | 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: 
69753diff
changeset | 1586 | 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: 
69753diff
changeset | 1587 | proof (rule connected_component_of_maximal) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1588 | 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: 
69753diff
changeset | 1589 | 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: 
69753diff
changeset | 1590 | 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: 
69753diff
changeset | 1591 | 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: 
69753diff
changeset | 1592 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1593 | ultimately | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1594 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1595 | 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: 
69753diff
changeset | 1596 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1597 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1598 | lemma closedin_connected_component_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1599 | "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: 
69753diff
changeset | 1600 | 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: 
69753diff
changeset | 1601 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1602 | lemma connected_component_of_eq_overlap: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1603 | "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: 
69753diff
changeset | 1604 | (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: 
69753diff
changeset | 1605 |       ~(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: 
69753diff
changeset | 1606 | using connected_component_of_equiv by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1607 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1608 | lemma connected_component_of_nonoverlap: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1609 |    "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: 
69753diff
changeset | 1610 | (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: 
69753diff
changeset | 1611 | ~(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: 
69753diff
changeset | 1612 | 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: 
69753diff
changeset | 1613 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1614 | lemma connected_component_of_overlap: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1615 |    "~(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: 
69753diff
changeset | 1616 | 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: 
69753diff
changeset | 1617 | 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: 
69753diff
changeset | 1618 | by (meson connected_component_of_nonoverlap) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69753diff
changeset | 1619 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1620 | 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: 
77934diff
changeset | 1621 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1622 | 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: 
77934diff
changeset | 1623 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1624 | lemma continuous_map_real_shrink: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1625 |   "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: 
77934diff
changeset | 1626 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1627 | 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: 
77934diff
changeset | 1628 | by (intro continuous_intros) auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1629 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1630 | by (auto simp: continuous_map_in_subtopology divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1631 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1632 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1633 | lemma continuous_on_real_grow: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1634 |   "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: 
77934diff
changeset | 1635 | by (intro continuous_intros) auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1636 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1637 | lemma real_grow_shrink: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1638 | fixes x::real | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1639 | 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: 
77934diff
changeset | 1640 | by (force simp: divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1641 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1642 | lemma homeomorphic_maps_real_shrink: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1643 |   "homeomorphic_maps euclideanreal (subtopology euclideanreal {-1<..<1}) 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77934diff
changeset | 1644 | (\<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: 
77934diff
changeset | 1645 | 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: 
77934diff
changeset | 1646 | |
| 78037 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1647 | lemma real_shrink_Galois: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1648 | fixes x::real | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1649 | 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: 
77943diff
changeset | 1650 | 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: 
77943diff
changeset | 1651 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 1652 | lemma real_shrink_eq: | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 1653 | fixes x y::real | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 1654 | 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: 
78037diff
changeset | 1655 | by (metis real_shrink_Galois) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 1656 | |
| 78037 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1657 | lemma real_shrink_lt: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1658 | fixes x::real | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1659 | 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: 
77943diff
changeset | 1660 | 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: 
77943diff
changeset | 1661 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1662 | lemma real_shrink_le: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1663 | fixes x::real | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1664 | 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: 
77943diff
changeset | 1665 | by (meson linorder_not_le real_shrink_lt) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1666 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1667 | lemma real_shrink_grow: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1668 | fixes x::real | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1669 | 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: 
77943diff
changeset | 1670 | using real_shrink_Galois by blast | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1671 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1672 | lemma continuous_shrink: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1673 | "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: 
77943diff
changeset | 1674 | by (intro continuous_intros) auto | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1675 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1676 | lemma strict_mono_shrink: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1677 | "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: 
77943diff
changeset | 1678 | by (simp add: monotoneI real_shrink_lt) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1679 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1680 | 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: 
77943diff
changeset | 1681 | by (auto simp: divide_simps) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1682 | |
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1683 | 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: 
77943diff
changeset | 1684 | lemma connected_shrink: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1685 | fixes S :: "real set" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1686 | 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: 
77943diff
changeset | 1687 | proof | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1688 | assume "?lhs" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1689 | 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: 
77943diff
changeset | 1690 | 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: 
77943diff
changeset | 1691 | continuous_on_subset) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1692 | then show "?rhs" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1693 | 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: 
77943diff
changeset | 1694 | next | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1695 | assume ?rhs | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1696 | then show ?lhs | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1697 | using connected_continuous_image | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1698 | 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: 
77943diff
changeset | 1699 | qed | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1700 | |
| 78256 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1701 | subsection \<open>A few cardinality results\<close> | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1702 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1703 | lemma eqpoll_real_subset: | 
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1704 | fixes a b::real | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1705 | 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: 
78248diff
changeset | 1706 | shows "S \<approx> (UNIV::real set)" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1707 | proof (rule lepoll_antisym) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1708 | show "S \<lesssim> (UNIV::real set)" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1709 | by (simp add: subset_imp_lepoll) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1710 | 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: 
78248diff
changeset | 1711 | show "(UNIV::real set) \<lesssim> S" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1712 | unfolding lepoll_def | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1713 | proof (intro exI conjI) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1714 | show "inj f" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1715 | unfolding inj_on_def f_def | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1716 | 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: 
78248diff
changeset | 1717 | have pos: "(b-a) / 2 > 0" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1718 | 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: 
78248diff
changeset | 1719 | 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: 
78248diff
changeset | 1720 | "(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: 
78248diff
changeset | 1721 | by (auto simp: field_simps) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1722 | show "range f \<subseteq> S" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1723 | 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: 
78248diff
changeset | 1724 | 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: 
78248diff
changeset | 1725 | 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: 
78248diff
changeset | 1726 | qed | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1727 | qed | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1728 | |
| 78256 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1729 | 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: 
78250diff
changeset | 1730 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1731 | 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: 
78250diff
changeset | 1732 |   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: 
78250diff
changeset | 1733 | by (simp add: nxt_def Pi_iff) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1734 | 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: 
78250diff
changeset | 1735 | 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: 
78250diff
changeset | 1736 | by (simp add: \<sigma>_def case_prod_beta') | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1737 |   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: 
78250diff
changeset | 1738 | proof (induction n) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1739 | case 0 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1740 | then show ?case | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1741 | by (simp add: \<sigma>_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1742 | next | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1743 | case (Suc n) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1744 | with nxt_fun show ?case | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1745 | by (force simp add: Pi_iff split: prod.split) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1746 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1747 |   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: 
78250diff
changeset | 1748 | 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: 
78250diff
changeset | 1749 | if "fst (nxt x) = fst (nxt y)" for x y | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1750 | 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: 
78250diff
changeset | 1751 | 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: 
78250diff
changeset | 1752 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1753 | 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: 
78250diff
changeset | 1754 | using that by (force simp add: f_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1755 | 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: 
78250diff
changeset | 1756 | 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: 
78250diff
changeset | 1757 | 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: 
78250diff
changeset | 1758 | proof (induction k) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1759 | case 0 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1760 | then show ?case | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1761 | by (simp add: \<sigma>_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1762 | next | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1763 | case (Suc k) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1764 | then show ?case | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1765 | by (simp add: eq snd_nxt) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1766 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1767 | 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: 
78250diff
changeset | 1768 | have "2^n \<ge> 1 / (y - x)" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1769 | by (simp add: n_def power_of_nat_log_ge) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1770 | then have "2^n * (y-x) \<ge> 1" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1771 | 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: 
78250diff
changeset | 1772 | 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: 
78250diff
changeset | 1773 | by presburger | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1774 |     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: 
78250diff
changeset | 1775 | 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: 
78250diff
changeset | 1776 | ultimately show False by simp | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1777 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1778 |   then have "inj_on f {0..<1}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1779 | by (meson atLeastLessThan_iff linorder_inj_onI') | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1780 | then show ?thesis | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1781 | unfolding lepoll_def by blast | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1782 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1783 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1784 | 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: 
78250diff
changeset | 1785 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1786 | 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: 
78250diff
changeset | 1787 | have Fge0: "F S i \<ge> 0" for S i | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1788 | by (simp add: F_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1789 | have F: "summable (F S)" for S | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1790 | 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: 
78250diff
changeset | 1791 |   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: 
78250diff
changeset | 1792 | proof (cases n) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1793 | case (Suc n') | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1794 |     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: 
78250diff
changeset | 1795 | by (simp add: F_def sum_mono) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1796 | also have "\<dots> = (\<Sum>i=0..n'. inverse 3 ^ i)" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1797 | using Suc atLeast0AtMost lessThan_Suc_atMost by presburger | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1798 | also have "\<dots> = (3/2) * (1 - inverse 3 ^ n)" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1799 | 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: 
78250diff
changeset | 1800 | also have "\<dots> \<le> 3/2" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1801 | by (simp add: field_simps) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1802 | finally show ?thesis . | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1803 | qed auto | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1804 | 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: 
78250diff
changeset | 1805 | using F suminf_le_const by blast | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1806 | 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: 
78250diff
changeset | 1807 | 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: 
78250diff
changeset | 1808 | using F_def that by auto | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1809 | 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: 
78250diff
changeset | 1810 | using that F by (simp add: f_def suminf_le) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1811 |   have "f S \<in> {0..<1::real}" for S
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1812 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1813 | have "0 \<le> suminf (F S)" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1814 | using F by (simp add: F_def suminf_nonneg) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1815 | with F32[of S] show ?thesis | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1816 | by (auto simp: f_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1817 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1818 | moreover have "inj f" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1819 | proof | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1820 | fix S T | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1821 | assume "f S = f T" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1822 | show "S = T" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1823 | proof (rule ccontr) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1824 | assume "S \<noteq> T" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1825 |       then have ST_ne: "(S-T) \<union> (T-S) \<noteq> {}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1826 | by blast | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1827 | 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: 
78250diff
changeset | 1828 |       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: 
78250diff
changeset | 1829 | by (metis F add.commute suminf_split_initial_segment) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1830 |       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: 
78250diff
changeset | 1831 | if "n \<in> U" for U | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1832 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1833 | have "0 \<le> (\<Sum>k. F U (k + Suc n))" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1834 | by (metis F Fge0 suminf_nonneg summable_iff_shift) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1835 | moreover have "F U n = (1/3) ^ n" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1836 | by (simp add: F_def that) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1837 | ultimately show ?thesis | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1838 | by (simp add: sum_split f_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1839 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1840 | 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: 
78250diff
changeset | 1841 | by (simp add: F_def power_add suminf_mult2) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1842 |       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: 
78250diff
changeset | 1843 | if "n \<notin> U" for U | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1844 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1845 | have [simp]: "F U n = 0" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1846 | by (simp add: F_def that) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1847 | 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: 
78250diff
changeset | 1848 | 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: 
78250diff
changeset | 1849 | 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: 
78250diff
changeset | 1850 | by (simp add: sum_split) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1851 | 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: 
78250diff
changeset | 1852 | unfolding * using F32[of UNIV] by simp | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1853 |         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: 
78250diff
changeset | 1854 | then show ?thesis | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1855 | by (simp add: f_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1856 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1857 |       have "S \<inter> {..<n} = T \<inter> {..<n}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1858 | using not_less_Least by (fastforce simp add: n_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1859 |       then have "sum (F S) {..<n} = sum (F T) {..<n}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1860 | 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: 
78250diff
changeset | 1861 | moreover consider "n \<in> S-T" | "n \<in> T-S" | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1862 | 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: 
78250diff
changeset | 1863 | ultimately show False | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1864 | 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: 
78250diff
changeset | 1865 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1866 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1867 | ultimately show ?thesis | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1868 | by (meson image_subsetI lepoll_def) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1869 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1870 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1871 | lemma open_interval_eqpoll_reals: | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1872 | fixes a b::real | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1873 |   shows "{a<..<b} \<approx> (UNIV::real set) \<longleftrightarrow> a<b"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1874 | using bij_betw_tan bij_betw_open_intervals eqpoll_def | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1875 | 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: 
78250diff
changeset | 1876 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1877 | lemma closed_interval_eqpoll_reals: | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1878 | fixes a b::real | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1879 |   shows "{a..b} \<approx> (UNIV::real set) \<longleftrightarrow> a < b"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1880 | proof | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1881 |   show "{a..b} \<approx> (UNIV::real set) \<Longrightarrow> a < b"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1882 | 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: 
78250diff
changeset | 1883 | qed (auto simp: eqpoll_real_subset) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1884 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1885 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1886 | 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: 
78250diff
changeset | 1887 | proof - | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1888 |   have "(UNIV::real set) \<approx> {0<..<1::real}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1889 | by (simp add: open_interval_eqpoll_reals eqpoll_sym) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1890 |   also have "\<dots> \<lesssim> {0..<1::real}"
 | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1891 | by (simp add: greaterThanLessThan_subseteq_atLeastLessThan_iff subset_imp_lepoll) | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1892 | finally show ?thesis . | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1893 | qed | 
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1894 | |
| 
71e1aa0d9421
A couple of new lemmas involving cardinality
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 1895 | 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: 
78250diff
changeset | 1896 | 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: 
78250diff
changeset | 1897 | 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: 
78250diff
changeset | 1898 | |
| 69616 
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
 immler parents: diff
changeset | 1899 | end |