| author | haftmann |
| Sun, 10 Mar 2019 15:16:45 +0000 | |
| changeset 69906 | 55534affe445 |
| parent 69753 | 9a3b4cca6d0b |
| child 69922 | 4a9167f377b0 |
| 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 |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
11 |
Elementary_Topology |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
12 |
Abstract_Topology |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
13 |
"HOL-Library.Indicator_Function" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
14 |
begin |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
15 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
16 |
text \<open>Combination of Elementary and Abstract Topology\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
17 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
18 |
(* FIXME: move elsewhere *) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
19 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
20 |
lemma approachable_lt_le: "(\<exists>(d::real) > 0. \<forall>x. f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> P x)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
21 |
apply auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
22 |
apply (rule_tac x="d/2" in exI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
23 |
apply auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
24 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
25 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
26 |
lemma approachable_lt_le2: \<comment> \<open>like the above, but pushes aside an extra formula\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
27 |
"(\<exists>(d::real) > 0. \<forall>x. Q x \<longrightarrow> f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> Q x \<longrightarrow> P x)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
28 |
apply auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
29 |
apply (rule_tac x="d/2" in exI, auto) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
30 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
31 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
32 |
lemma triangle_lemma: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
33 |
fixes x y z :: real |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
34 |
assumes x: "0 \<le> x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
35 |
and y: "0 \<le> y" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
36 |
and z: "0 \<le> z" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
37 |
and xy: "x\<^sup>2 \<le> y\<^sup>2 + z\<^sup>2" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
38 |
shows "x \<le> y + z" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
39 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
40 |
have "y\<^sup>2 + z\<^sup>2 \<le> y\<^sup>2 + 2 * y * z + z\<^sup>2" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
41 |
using z y by simp |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
42 |
with xy have th: "x\<^sup>2 \<le> (y + z)\<^sup>2" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
43 |
by (simp add: power2_eq_square field_simps) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
44 |
from y z have yz: "y + z \<ge> 0" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
45 |
by arith |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
46 |
from power2_le_imp_le[OF th yz] show ?thesis . |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
47 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
48 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
49 |
lemma isCont_indicator: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
50 |
fixes x :: "'a::t2_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
51 |
shows "isCont (indicator A :: 'a \<Rightarrow> real) x = (x \<notin> frontier A)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
52 |
proof auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
53 |
fix x |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
54 |
assume cts_at: "isCont (indicator A :: 'a \<Rightarrow> real) x" and fr: "x \<in> frontier A" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
55 |
with continuous_at_open have 1: "\<forall>V::real set. open V \<and> indicator A x \<in> V \<longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
56 |
(\<exists>U::'a set. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> V))" by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
57 |
show False |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
58 |
proof (cases "x \<in> A") |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
59 |
assume x: "x \<in> A" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
60 |
hence "indicator A x \<in> ({0<..<2} :: real set)" by simp
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
61 |
hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({0<..<2} :: real set))"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
62 |
using 1 open_greaterThanLessThan by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
63 |
then guess U .. note U = this |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
64 |
hence "\<forall>y\<in>U. indicator A y > (0::real)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
65 |
unfolding greaterThanLessThan_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
66 |
hence "U \<subseteq> A" using indicator_eq_0_iff by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
67 |
hence "x \<in> interior A" using U interiorI by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
68 |
thus ?thesis using fr unfolding frontier_def by simp |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
69 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
70 |
assume x: "x \<notin> A" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
71 |
hence "indicator A x \<in> ({-1<..<1} :: real set)" by simp
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
72 |
hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({-1<..<1} :: real set))"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
73 |
using 1 open_greaterThanLessThan by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
74 |
then guess U .. note U = this |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
75 |
hence "\<forall>y\<in>U. indicator A y < (1::real)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
76 |
unfolding greaterThanLessThan_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
77 |
hence "U \<subseteq> -A" by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
78 |
hence "x \<in> interior (-A)" using U interiorI by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
79 |
thus ?thesis using fr interior_complement unfolding frontier_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
80 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
81 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
82 |
assume nfr: "x \<notin> frontier A" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
83 |
hence "x \<in> interior A \<or> x \<in> interior (-A)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
84 |
by (auto simp: frontier_def closure_interior) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
85 |
thus "isCont ((indicator A)::'a \<Rightarrow> real) x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
86 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
87 |
assume int: "x \<in> interior A" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
88 |
then obtain U where U: "open U" "x \<in> U" "U \<subseteq> A" unfolding interior_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
89 |
hence "\<forall>y\<in>U. indicator A y = (1::real)" unfolding indicator_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
90 |
hence "continuous_on U (indicator A)" by (simp add: continuous_on_const indicator_eq_1_iff) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
91 |
thus ?thesis using U continuous_on_eq_continuous_at by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
92 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
93 |
assume ext: "x \<in> interior (-A)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
94 |
then obtain U where U: "open U" "x \<in> U" "U \<subseteq> -A" unfolding interior_def by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
95 |
then have "continuous_on U (indicator A)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
96 |
using continuous_on_topological by (auto simp: subset_iff) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
97 |
thus ?thesis using U continuous_on_eq_continuous_at by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
98 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
99 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
100 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
101 |
lemma closedin_limpt: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
102 |
"closedin (subtopology euclidean T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> S)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
103 |
apply (simp add: closedin_closed, safe) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
104 |
apply (simp add: closed_limpt islimpt_subset) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
105 |
apply (rule_tac x="closure S" in exI, simp) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
106 |
apply (force simp: closure_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
107 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
108 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
109 |
lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (subtopology euclidean S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
110 |
by (meson closedin_limpt closed_subset closedin_closed_trans) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
111 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
112 |
lemma connected_closed_set: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
113 |
"closed S |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
114 |
\<Longrightarrow> connected S \<longleftrightarrow> (\<nexists>A B. closed A \<and> closed B \<and> A \<noteq> {} \<and> B \<noteq> {} \<and> A \<union> B = S \<and> A \<inter> B = {})"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
115 |
unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
116 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
117 |
text \<open>If a connnected set is written as the union of two nonempty closed sets, then these sets |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
118 |
have to intersect.\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
119 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
120 |
lemma connected_as_closed_union: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
121 |
assumes "connected C" "C = A \<union> B" "closed A" "closed B" "A \<noteq> {}" "B \<noteq> {}"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
122 |
shows "A \<inter> B \<noteq> {}"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
123 |
by (metis assms closed_Un connected_closed_set) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
124 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
125 |
lemma closedin_subset_trans: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
126 |
"closedin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
127 |
closedin (subtopology euclidean T) S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
128 |
by (meson closedin_limpt subset_iff) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
129 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
130 |
lemma openin_subset_trans: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
131 |
"openin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
132 |
openin (subtopology euclidean T) S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
133 |
by (auto simp: openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
134 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
135 |
lemma closedin_compact: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
136 |
"\<lbrakk>compact S; closedin (subtopology euclidean S) T\<rbrakk> \<Longrightarrow> compact T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
137 |
by (metis closedin_closed compact_Int_closed) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
138 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
139 |
lemma closedin_compact_eq: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
140 |
fixes S :: "'a::t2_space set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
141 |
shows |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
142 |
"compact S |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
143 |
\<Longrightarrow> (closedin (subtopology euclidean S) T \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
144 |
compact T \<and> T \<subseteq> S)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
145 |
by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
146 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
147 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
148 |
subsection \<open>Closure\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
149 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
150 |
lemma closure_openin_Int_closure: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
151 |
assumes ope: "openin (subtopology euclidean U) S" and "T \<subseteq> U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
152 |
shows "closure(S \<inter> closure T) = closure(S \<inter> T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
153 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
154 |
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
|
155 |
using ope using openin_open by metis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
156 |
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
|
157 |
proof (clarsimp simp: S) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
158 |
fix x |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
159 |
assume "x \<in> closure (U \<inter> V \<inter> closure T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
160 |
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
|
161 |
by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
162 |
then have "x \<in> closure (T \<inter> V)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
163 |
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
|
164 |
then show "x \<in> closure (U \<inter> V \<inter> T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
165 |
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
|
166 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
167 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
168 |
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
|
169 |
by (meson Int_mono closure_mono closure_subset order_refl) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
170 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
171 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
172 |
corollary infinite_openin: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
173 |
fixes S :: "'a :: t1_space set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
174 |
shows "\<lbrakk>openin (subtopology euclidean U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
175 |
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
|
176 |
|
| 69622 | 177 |
lemma closure_Int_ballI: |
178 |
assumes "\<And>U. \<lbrakk>openin (subtopology euclidean S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
|
|
179 |
shows "S \<subseteq> closure T" |
|
180 |
proof (clarsimp simp: closure_iff_nhds_not_empty) |
|
181 |
fix x and A and V |
|
182 |
assume "x \<in> S" "V \<subseteq> A" "open V" "x \<in> V" "T \<inter> A = {}"
|
|
183 |
then have "openin (subtopology euclidean S) (A \<inter> V \<inter> S)" |
|
184 |
by (auto simp: openin_open intro!: exI[where x="V"]) |
|
185 |
moreover have "A \<inter> V \<inter> S \<noteq> {}" using \<open>x \<in> V\<close> \<open>V \<subseteq> A\<close> \<open>x \<in> S\<close>
|
|
186 |
by auto |
|
187 |
ultimately have "T \<inter> (A \<inter> V \<inter> S) \<noteq> {}"
|
|
188 |
by (rule assms) |
|
189 |
with \<open>T \<inter> A = {}\<close> show False by auto
|
|
190 |
qed |
|
191 |
||
192 |
||
|
69616
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
193 |
subsection \<open>Frontier\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
194 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
195 |
lemma connected_Int_frontier: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
196 |
"\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
197 |
apply (simp add: frontier_interiors connected_openin, safe) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
198 |
apply (drule_tac x="s \<inter> interior t" in spec, safe) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
199 |
apply (drule_tac [2] x="s \<inter> interior (-t)" in spec) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
200 |
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
|
201 |
done |
|
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 |
subsection \<open>Compactness\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
204 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
205 |
lemma openin_delete: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
206 |
fixes a :: "'a :: t1_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
207 |
shows "openin (subtopology euclidean u) s |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
208 |
\<Longrightarrow> openin (subtopology euclidean u) (s - {a})"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
209 |
by (metis Int_Diff open_delete openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
210 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
211 |
lemma compact_eq_openin_cover: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
212 |
"compact S \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
213 |
(\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
214 |
(\<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
|
215 |
proof safe |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
216 |
fix C |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
217 |
assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
218 |
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
|
219 |
unfolding openin_open by force+ |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
220 |
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
|
221 |
by (meson compactE) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
222 |
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
|
223 |
by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
224 |
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
|
225 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
226 |
assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
227 |
(\<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
|
228 |
show "compact S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
229 |
proof (rule compactI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
230 |
fix C |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
231 |
let ?C = "image (\<lambda>T. S \<inter> T) C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
232 |
assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
233 |
then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
234 |
unfolding openin_open by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
235 |
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
|
236 |
by metis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
237 |
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
|
238 |
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
|
239 |
proof (intro conjI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
240 |
from \<open>D \<subseteq> ?C\<close> show "?D \<subseteq> C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
241 |
by (fast intro: inv_into_into) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
242 |
from \<open>finite D\<close> show "finite ?D" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
243 |
by (rule finite_imageI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
244 |
from \<open>S \<subseteq> \<Union>D\<close> show "S \<subseteq> \<Union>?D" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
245 |
apply (rule subset_trans, clarsimp) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
246 |
apply (frule subsetD [OF \<open>D \<subseteq> ?C\<close>, THEN f_inv_into_f]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
247 |
apply (erule rev_bexI, fast) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
248 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
249 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
250 |
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
|
251 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
252 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
253 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
254 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
255 |
subsection \<open>Continuity\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
256 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
257 |
lemma interior_image_subset: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
258 |
assumes "inj f" "\<And>x. continuous (at x) f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
259 |
shows "interior (f ` S) \<subseteq> f ` (interior S)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
260 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
261 |
fix x assume "x \<in> interior (f ` S)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
262 |
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
|
263 |
then have "x \<in> f ` S" by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
264 |
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
|
265 |
have "open (f -` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
266 |
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
|
267 |
moreover have "y \<in> vimage f T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
268 |
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
|
269 |
moreover have "vimage f T \<subseteq> S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
270 |
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
|
271 |
ultimately have "y \<in> interior S" .. |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
272 |
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
|
273 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
274 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
275 |
subsection%unimportant \<open>Equality of continuous functions on closure and related results\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
276 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
277 |
lemma continuous_closedin_preimage_constant: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
278 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
279 |
shows "continuous_on S f \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x = a}"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
280 |
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
|
281 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
282 |
lemma continuous_closed_preimage_constant: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
283 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
284 |
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
|
285 |
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
|
286 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
287 |
lemma continuous_constant_on_closure: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
288 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
289 |
assumes "continuous_on (closure S) f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
290 |
and "\<And>x. x \<in> S \<Longrightarrow> f x = a" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
291 |
and "x \<in> closure S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
292 |
shows "f x = a" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
293 |
using continuous_closed_preimage_constant[of "closure S" f a] |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
294 |
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
|
295 |
unfolding subset_eq |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
296 |
by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
297 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
298 |
lemma image_closure_subset: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
299 |
assumes contf: "continuous_on (closure S) f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
300 |
and "closed T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
301 |
and "(f ` S) \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
302 |
shows "f ` (closure S) \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
303 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
304 |
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
|
305 |
using assms(3) closure_subset by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
306 |
moreover have "closed (closure S \<inter> f -` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
307 |
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
|
308 |
ultimately have "closure S = (closure S \<inter> f -` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
309 |
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
|
310 |
then show ?thesis by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
311 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
312 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
313 |
subsection%unimportant \<open>A function constant on a set\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
314 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
315 |
definition constant_on (infixl "(constant'_on)" 50) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
316 |
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
|
317 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
318 |
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
|
319 |
unfolding constant_on_def by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
320 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
321 |
lemma injective_not_constant: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
322 |
fixes S :: "'a::{perfect_space} set"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
323 |
shows "\<lbrakk>open S; inj_on f S; f constant_on S\<rbrakk> \<Longrightarrow> S = {}"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
324 |
unfolding constant_on_def |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
325 |
by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
326 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
327 |
lemma constant_on_closureI: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
328 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
329 |
assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
330 |
shows "f constant_on (closure S)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
331 |
using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
332 |
by metis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
333 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
334 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
335 |
subsection%unimportant \<open>Continuity relative to a union.\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
336 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
337 |
lemma continuous_on_Un_local: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
338 |
"\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t; |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
339 |
continuous_on s f; continuous_on t f\<rbrakk> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
340 |
\<Longrightarrow> continuous_on (s \<union> t) f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
341 |
unfolding continuous_on closedin_limpt |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
342 |
by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
343 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
344 |
lemma continuous_on_cases_local: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
345 |
"\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t; |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
346 |
continuous_on s f; continuous_on t g; |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
347 |
\<And>x. \<lbrakk>x \<in> s \<and> \<not>P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
348 |
\<Longrightarrow> continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
349 |
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
|
350 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
351 |
lemma continuous_on_cases_le: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
352 |
fixes h :: "'a :: topological_space \<Rightarrow> real" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
353 |
assumes "continuous_on {t \<in> s. h t \<le> a} f"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
354 |
and "continuous_on {t \<in> s. a \<le> h t} g"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
355 |
and h: "continuous_on s h" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
356 |
and "\<And>t. \<lbrakk>t \<in> s; h t = a\<rbrakk> \<Longrightarrow> f t = g t" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
357 |
shows "continuous_on s (\<lambda>t. if h t \<le> a then f(t) else g(t))" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
358 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
359 |
have s: "s = (s \<inter> h -` atMost a) \<union> (s \<inter> h -` atLeast a)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
360 |
by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
361 |
have 1: "closedin (subtopology euclidean s) (s \<inter> h -` atMost a)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
362 |
by (rule continuous_closedin_preimage [OF h closed_atMost]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
363 |
have 2: "closedin (subtopology euclidean s) (s \<inter> h -` atLeast a)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
364 |
by (rule continuous_closedin_preimage [OF h closed_atLeast]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
365 |
have eq: "s \<inter> h -` {..a} = {t \<in> s. h t \<le> a}" "s \<inter> h -` {a..} = {t \<in> s. a \<le> h t}"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
366 |
by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
367 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
368 |
apply (rule continuous_on_subset [of s, OF _ order_refl]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
369 |
apply (subst s) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
370 |
apply (rule continuous_on_cases_local) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
371 |
using 1 2 s assms apply (auto simp: eq) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
372 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
373 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
374 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
375 |
lemma continuous_on_cases_1: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
376 |
fixes s :: "real set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
377 |
assumes "continuous_on {t \<in> s. t \<le> a} f"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
378 |
and "continuous_on {t \<in> s. a \<le> t} g"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
379 |
and "a \<in> s \<Longrightarrow> f a = g a" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
380 |
shows "continuous_on s (\<lambda>t. if t \<le> a then f(t) else g(t))" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
381 |
using assms |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
382 |
by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
383 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
384 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
385 |
subsection%unimportant\<open>Inverse function property for open/closed maps\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
386 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
387 |
lemma continuous_on_inverse_open_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
388 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
389 |
and imf: "f ` S = T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
390 |
and injf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
391 |
and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
392 |
shows "continuous_on T g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
393 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
394 |
from imf injf have gTS: "g ` T = S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
395 |
by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
396 |
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
|
397 |
by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
398 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
399 |
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
|
400 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
401 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
402 |
lemma continuous_on_inverse_closed_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
403 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
404 |
and imf: "f ` S = T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
405 |
and injf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
406 |
and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
407 |
shows "continuous_on T g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
408 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
409 |
from imf injf have gTS: "g ` T = S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
410 |
by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
411 |
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
|
412 |
by force |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
413 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
414 |
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
|
415 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
416 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
417 |
lemma homeomorphism_injective_open_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
418 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
419 |
and imf: "f ` S = T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
420 |
and injf: "inj_on f S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
421 |
and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
422 |
obtains g where "homeomorphism S T f g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
423 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
424 |
have "continuous_on T (inv_into S f)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
425 |
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
|
426 |
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
|
427 |
by (auto simp: homeomorphism_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
428 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
429 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
430 |
lemma homeomorphism_injective_closed_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
431 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
432 |
and imf: "f ` S = T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
433 |
and injf: "inj_on f S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
434 |
and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
435 |
obtains g where "homeomorphism S T f g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
436 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
437 |
have "continuous_on T (inv_into S f)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
438 |
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
|
439 |
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
|
440 |
by (auto simp: homeomorphism_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
441 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
442 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
443 |
lemma homeomorphism_imp_open_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
444 |
assumes hom: "homeomorphism S T f g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
445 |
and oo: "openin (subtopology euclidean S) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
446 |
shows "openin (subtopology euclidean T) (f ` U)" |
|
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 |
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
|
449 |
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
|
450 |
from hom have "continuous_on T g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
451 |
unfolding homeomorphism_def by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
452 |
moreover have "g ` T = S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
453 |
by (metis hom homeomorphism_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
454 |
ultimately show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
455 |
by (simp add: continuous_on_open oo) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
456 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
457 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
458 |
lemma homeomorphism_imp_closed_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
459 |
assumes hom: "homeomorphism S T f g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
460 |
and oo: "closedin (subtopology euclidean S) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
461 |
shows "closedin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
462 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
463 |
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
|
464 |
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
|
465 |
from hom have "continuous_on T g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
466 |
unfolding homeomorphism_def by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
467 |
moreover have "g ` T = S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
468 |
by (metis hom homeomorphism_def) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
469 |
ultimately show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
470 |
by (simp add: continuous_on_closed oo) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
471 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
472 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
473 |
subsection%unimportant \<open>Seperability\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
474 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
475 |
lemma subset_second_countable: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
476 |
obtains \<B> :: "'a:: second_countable_topology set set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
477 |
where "countable \<B>" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
478 |
"{} \<notin> \<B>"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
479 |
"\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
480 |
"\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
481 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
482 |
obtain \<B> :: "'a set set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
483 |
where "countable \<B>" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
484 |
and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
485 |
and \<B>: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
486 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
487 |
obtain \<C> :: "'a set set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
488 |
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
|
489 |
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
|
490 |
by (metis univ_second_countable that) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
491 |
show ?thesis |
|
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 |
show "countable ((\<lambda>C. S \<inter> C) ` \<C>)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
494 |
by (simp add: \<open>countable \<C>\<close>) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
495 |
show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (subtopology euclidean S) C" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
496 |
using ope by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
497 |
show "\<And>T. openin (subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
498 |
by (metis \<C> image_mono inf_Sup openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
499 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
500 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
501 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
502 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
503 |
show "countable (\<B> - {{}})"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
504 |
using \<open>countable \<B>\<close> by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
505 |
show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) C"
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
506 |
by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (subtopology euclidean S) C\<close>) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
507 |
show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (subtopology euclidean S) T" for T
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
508 |
using \<B> [OF that] |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
509 |
apply clarify |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
510 |
apply (rule_tac x="\<U> - {{}}" in exI, auto)
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
511 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
512 |
qed auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
513 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
514 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
515 |
lemma Lindelof_openin: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
516 |
fixes \<F> :: "'a::second_countable_topology set set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
517 |
assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean U) S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
518 |
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
|
519 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
520 |
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
|
521 |
using assms by (simp add: openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
522 |
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
|
523 |
by metis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
524 |
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
|
525 |
using tf by fastforce |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
526 |
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
|
527 |
using tf by (force intro: Lindelof [of "tf ` \<F>"]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
528 |
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
|
529 |
by (clarsimp simp add: countable_subset_image) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
530 |
then show ?thesis .. |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
531 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
532 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
533 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
534 |
subsection%unimportant\<open>Closed Maps\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
535 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
536 |
lemma continuous_imp_closed_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
537 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
538 |
assumes "closedin (subtopology euclidean S) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
539 |
"continuous_on S f" "f ` S = T" "compact S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
540 |
shows "closedin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
541 |
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
|
542 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
543 |
lemma closed_map_restrict: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
544 |
assumes cloU: "closedin (subtopology euclidean (S \<inter> f -` T')) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
545 |
and cc: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
546 |
and "T' \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
547 |
shows "closedin (subtopology euclidean T') (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
548 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
549 |
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
|
550 |
using cloU by (auto simp: closedin_closed) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
551 |
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
|
552 |
by (fastforce simp add: closedin_closed) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
553 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
554 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
555 |
subsection%unimportant\<open>Open Maps\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
556 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
557 |
lemma open_map_restrict: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
558 |
assumes opeU: "openin (subtopology euclidean (S \<inter> f -` T')) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
559 |
and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
560 |
and "T' \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
561 |
shows "openin (subtopology euclidean T') (f ` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
562 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
563 |
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
|
564 |
using opeU by (auto simp: openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
565 |
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
|
566 |
by (fastforce simp add: openin_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
567 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
568 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
569 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
570 |
subsection%unimportant\<open>Quotient maps\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
571 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
572 |
lemma quotient_map_imp_continuous_open: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
573 |
assumes T: "f ` S \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
574 |
and ope: "\<And>U. U \<subseteq> T |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
575 |
\<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
576 |
openin (subtopology euclidean T) U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
577 |
shows "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
578 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
579 |
have [simp]: "S \<inter> f -` f ` S = S" by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
580 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
581 |
using ope [OF T] |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
582 |
apply (simp add: continuous_on_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
583 |
by (meson ope openin_imp_subset openin_trans) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
584 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
585 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
586 |
lemma quotient_map_imp_continuous_closed: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
587 |
assumes T: "f ` S \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
588 |
and ope: "\<And>U. U \<subseteq> T |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
589 |
\<Longrightarrow> (closedin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
590 |
closedin (subtopology euclidean T) U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
591 |
shows "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
592 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
593 |
have [simp]: "S \<inter> f -` f ` S = S" by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
594 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
595 |
using ope [OF T] |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
596 |
apply (simp add: continuous_on_closed) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
597 |
by (metis (no_types, lifting) ope closedin_imp_subset closedin_trans) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
598 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
599 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
600 |
lemma open_map_imp_quotient_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
601 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
602 |
and T: "T \<subseteq> f ` S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
603 |
and ope: "\<And>T. openin (subtopology euclidean S) T |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
604 |
\<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
605 |
shows "openin (subtopology euclidean S) (S \<inter> f -` T) = |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
606 |
openin (subtopology euclidean (f ` S)) T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
607 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
608 |
have "T = f ` (S \<inter> f -` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
609 |
using T by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
610 |
then show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
611 |
using "ope" contf continuous_on_open by metis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
612 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
613 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
614 |
lemma closed_map_imp_quotient_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
615 |
assumes contf: "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
616 |
and T: "T \<subseteq> f ` S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
617 |
and ope: "\<And>T. closedin (subtopology euclidean S) T |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
618 |
\<Longrightarrow> closedin (subtopology euclidean (f ` S)) (f ` T)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
619 |
shows "openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
620 |
openin (subtopology euclidean (f ` S)) T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
621 |
(is "?lhs = ?rhs") |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
622 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
623 |
assume ?lhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
624 |
then have *: "closedin (subtopology euclidean S) (S - (S \<inter> f -` T))" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
625 |
using closedin_diff by fastforce |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
626 |
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
|
627 |
using T by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
628 |
show ?rhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
629 |
using ope [OF *, unfolded closedin_def] by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
630 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
631 |
assume ?rhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
632 |
with contf show ?lhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
633 |
by (auto simp: continuous_on_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
634 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
635 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
636 |
lemma continuous_right_inverse_imp_quotient_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
637 |
assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
638 |
and contg: "continuous_on T g" and img: "g ` T \<subseteq> S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
639 |
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
|
640 |
and U: "U \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
641 |
shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
642 |
openin (subtopology euclidean T) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
643 |
(is "?lhs = ?rhs") |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
644 |
proof - |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
645 |
have f: "\<And>Z. openin (subtopology euclidean (f ` S)) Z \<Longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
646 |
openin (subtopology euclidean S) (S \<inter> f -` Z)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
647 |
and g: "\<And>Z. openin (subtopology euclidean (g ` T)) Z \<Longrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
648 |
openin (subtopology euclidean T) (T \<inter> g -` Z)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
649 |
using contf contg by (auto simp: continuous_on_open) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
650 |
show ?thesis |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
651 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
652 |
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
|
653 |
using imf img by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
654 |
also have "... = U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
655 |
using U by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
656 |
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
|
657 |
assume ?lhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
658 |
then have *: "openin (subtopology euclidean (g ` T)) (g ` T \<inter> (S \<inter> f -` U))" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
659 |
by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
660 |
show ?rhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
661 |
using g [OF *] eq by auto |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
662 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
663 |
assume rhs: ?rhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
664 |
show ?lhs |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
665 |
by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
666 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
667 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
668 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
669 |
lemma continuous_left_inverse_imp_quotient_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
670 |
assumes "continuous_on S f" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
671 |
and "continuous_on (f ` S) g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
672 |
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
|
673 |
and "U \<subseteq> f ` S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
674 |
shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
675 |
openin (subtopology euclidean (f ` S)) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
676 |
apply (rule continuous_right_inverse_imp_quotient_map) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
677 |
using assms apply force+ |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
678 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
679 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
680 |
lemma continuous_imp_quotient_map: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
681 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
682 |
assumes "continuous_on S f" "f ` S = T" "compact S" "U \<subseteq> T" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
683 |
shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
684 |
openin (subtopology euclidean T) U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
685 |
by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
686 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
687 |
subsection%unimportant\<open>Pasting functions together\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
688 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
689 |
text\<open>on open sets\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
690 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
691 |
lemma pasting_lemma: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
692 |
fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
693 |
assumes clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
694 |
and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
695 |
and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
696 |
and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
697 |
shows "continuous_on S g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
698 |
proof (clarsimp simp: continuous_openin_preimage_eq) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
699 |
fix U :: "'b set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
700 |
assume "open U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
701 |
have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
702 |
using clo openin_imp_subset by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
703 |
have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
704 |
using S f g by fastforce |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
705 |
show "openin (subtopology euclidean S) (S \<inter> g -` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
706 |
apply (subst *) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
707 |
apply (rule openin_Union, clarify) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
708 |
using \<open>open U\<close> clo cont continuous_openin_preimage_gen openin_trans by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
709 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
710 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
711 |
lemma pasting_lemma_exists: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
712 |
fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
713 |
assumes S: "S \<subseteq> (\<Union>i \<in> I. T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
714 |
and clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
715 |
and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
716 |
and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
717 |
obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
718 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
719 |
show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
720 |
apply (rule pasting_lemma [OF clo cont]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
721 |
apply (blast intro: f)+ |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
722 |
apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
723 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
724 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
725 |
fix x i |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
726 |
assume "i \<in> I" "x \<in> S \<inter> T i" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
727 |
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
|
728 |
by (metis (no_types, lifting) IntD2 IntI f someI_ex) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
729 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
730 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
731 |
text\<open>Likewise on closed sets, with a finiteness assumption\<close> |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
732 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
733 |
lemma pasting_lemma_closed: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
734 |
fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
735 |
assumes "finite I" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
736 |
and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
737 |
and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
738 |
and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
739 |
and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
740 |
shows "continuous_on S g" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
741 |
proof (clarsimp simp: continuous_closedin_preimage_eq) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
742 |
fix U :: "'b set" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
743 |
assume "closed U" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
744 |
have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
745 |
using clo closedin_imp_subset by blast |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
746 |
have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
747 |
using S f g by fastforce |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
748 |
show "closedin (subtopology euclidean S) (S \<inter> g -` U)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
749 |
apply (subst *) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
750 |
apply (rule closedin_Union) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
751 |
using \<open>finite I\<close> apply simp |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
752 |
apply (blast intro: \<open>closed U\<close> continuous_closedin_preimage cont clo closedin_trans) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
753 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
754 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
755 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
756 |
lemma pasting_lemma_exists_closed: |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
757 |
fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
758 |
assumes "finite I" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
759 |
and S: "S \<subseteq> (\<Union>i \<in> I. T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
760 |
and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
761 |
and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
762 |
and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
763 |
obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
764 |
proof |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
765 |
show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
766 |
apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont]) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
767 |
apply (blast intro: f)+ |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
768 |
apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
769 |
done |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
770 |
next |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
771 |
fix x i |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
772 |
assume "i \<in> I" "x \<in> S \<inter> T i" |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
773 |
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
|
774 |
by (metis (no_types, lifting) IntD2 IntI f someI_ex) |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
775 |
qed |
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
776 |
|
|
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
777 |
|
| 69750 | 778 |
subsection \<open>Retractions\<close> |
779 |
||
780 |
definition%important retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
|
|
781 |
where "retraction S T r \<longleftrightarrow> |
|
782 |
T \<subseteq> S \<and> continuous_on S r \<and> r ` S \<subseteq> T \<and> (\<forall>x\<in>T. r x = x)" |
|
783 |
||
784 |
definition%important retract_of (infixl "retract'_of" 50) where |
|
785 |
"T retract_of S \<longleftrightarrow> (\<exists>r. retraction S T r)" |
|
786 |
||
787 |
lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow> r (r x) = r x" |
|
788 |
unfolding retraction_def by auto |
|
789 |
||
790 |
text \<open>Preservation of fixpoints under (more general notion of) retraction\<close> |
|
791 |
||
792 |
lemma invertible_fixpoint_property: |
|
793 |
fixes S :: "'a::topological_space set" |
|
794 |
and T :: "'b::topological_space set" |
|
795 |
assumes contt: "continuous_on T i" |
|
796 |
and "i ` T \<subseteq> S" |
|
797 |
and contr: "continuous_on S r" |
|
798 |
and "r ` S \<subseteq> T" |
|
799 |
and ri: "\<And>y. y \<in> T \<Longrightarrow> r (i y) = y" |
|
800 |
and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" |
|
801 |
and contg: "continuous_on T g" |
|
802 |
and "g ` T \<subseteq> T" |
|
803 |
obtains y where "y \<in> T" and "g y = y" |
|
804 |
proof - |
|
805 |
have "\<exists>x\<in>S. (i \<circ> g \<circ> r) x = x" |
|
806 |
proof (rule FP) |
|
807 |
show "continuous_on S (i \<circ> g \<circ> r)" |
|
808 |
by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) |
|
809 |
show "(i \<circ> g \<circ> r) ` S \<subseteq> S" |
|
810 |
using assms(2,4,8) by force |
|
811 |
qed |
|
812 |
then obtain x where x: "x \<in> S" "(i \<circ> g \<circ> r) x = x" .. |
|
813 |
then have *: "g (r x) \<in> T" |
|
814 |
using assms(4,8) by auto |
|
815 |
have "r ((i \<circ> g \<circ> r) x) = r x" |
|
816 |
using x by auto |
|
817 |
then show ?thesis |
|
818 |
using "*" ri that by auto |
|
819 |
qed |
|
820 |
||
821 |
lemma homeomorphic_fixpoint_property: |
|
822 |
fixes S :: "'a::topological_space set" |
|
823 |
and T :: "'b::topological_space set" |
|
824 |
assumes "S homeomorphic T" |
|
825 |
shows "(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> S \<longrightarrow> (\<exists>x\<in>S. f x = x)) \<longleftrightarrow> |
|
826 |
(\<forall>g. continuous_on T g \<and> g ` T \<subseteq> T \<longrightarrow> (\<exists>y\<in>T. g y = y))" |
|
827 |
(is "?lhs = ?rhs") |
|
828 |
proof - |
|
829 |
obtain r i where r: |
|
830 |
"\<forall>x\<in>S. i (r x) = x" "r ` S = T" "continuous_on S r" |
|
831 |
"\<forall>y\<in>T. r (i y) = y" "i ` T = S" "continuous_on T i" |
|
832 |
using assms unfolding homeomorphic_def homeomorphism_def by blast |
|
833 |
show ?thesis |
|
834 |
proof |
|
835 |
assume ?lhs |
|
836 |
with r show ?rhs |
|
837 |
by (metis invertible_fixpoint_property[of T i S r] order_refl) |
|
838 |
next |
|
839 |
assume ?rhs |
|
840 |
with r show ?lhs |
|
841 |
by (metis invertible_fixpoint_property[of S r T i] order_refl) |
|
842 |
qed |
|
843 |
qed |
|
844 |
||
845 |
lemma retract_fixpoint_property: |
|
846 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
847 |
and S :: "'a set" |
|
848 |
assumes "T retract_of S" |
|
849 |
and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" |
|
850 |
and contg: "continuous_on T g" |
|
851 |
and "g ` T \<subseteq> T" |
|
852 |
obtains y where "y \<in> T" and "g y = y" |
|
853 |
proof - |
|
854 |
obtain h where "retraction S T h" |
|
855 |
using assms(1) unfolding retract_of_def .. |
|
856 |
then show ?thesis |
|
857 |
unfolding retraction_def |
|
858 |
using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] |
|
859 |
by (metis assms(4) contg image_ident that) |
|
860 |
qed |
|
861 |
||
862 |
lemma retraction: |
|
863 |
"retraction S T r \<longleftrightarrow> |
|
864 |
T \<subseteq> S \<and> continuous_on S r \<and> r ` S = T \<and> (\<forall>x \<in> T. r x = x)" |
|
865 |
by (force simp: retraction_def) |
|
866 |
||
| 69753 | 867 |
lemma retractionE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close> |
| 69750 | 868 |
assumes "retraction S T r" |
869 |
obtains "T = r ` S" "r ` S \<subseteq> S" "continuous_on S r" "\<And>x. x \<in> S \<Longrightarrow> r (r x) = r x" |
|
870 |
proof (rule that) |
|
871 |
from retraction [of S T r] assms |
|
872 |
have "T \<subseteq> S" "continuous_on S r" "r ` S = T" and "\<forall>x \<in> T. r x = x" |
|
873 |
by simp_all |
|
874 |
then show "T = r ` S" "r ` S \<subseteq> S" "continuous_on S r" |
|
875 |
by simp_all |
|
876 |
from \<open>\<forall>x \<in> T. r x = x\<close> have "r x = x" if "x \<in> T" for x |
|
877 |
using that by simp |
|
878 |
with \<open>r ` S = T\<close> show "r (r x) = r x" if "x \<in> S" for x |
|
879 |
using that by auto |
|
880 |
qed |
|
881 |
||
| 69753 | 882 |
lemma retract_ofE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close> |
| 69750 | 883 |
assumes "T retract_of S" |
884 |
obtains r where "T = r ` S" "r ` S \<subseteq> S" "continuous_on S r" "\<And>x. x \<in> S \<Longrightarrow> r (r x) = r x" |
|
885 |
proof - |
|
886 |
from assms obtain r where "retraction S T r" |
|
887 |
by (auto simp add: retract_of_def) |
|
888 |
with that show thesis |
|
889 |
by (auto elim: retractionE) |
|
890 |
qed |
|
891 |
||
892 |
lemma retract_of_imp_extensible: |
|
893 |
assumes "S retract_of T" and "continuous_on S f" and "f ` S \<subseteq> U" |
|
894 |
obtains g where "continuous_on T g" "g ` T \<subseteq> U" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" |
|
895 |
proof - |
|
896 |
from \<open>S retract_of T\<close> obtain r where "retraction T S r" |
|
897 |
by (auto simp add: retract_of_def) |
|
898 |
show thesis |
|
899 |
by (rule that [of "f \<circ> r"]) |
|
900 |
(use \<open>continuous_on S f\<close> \<open>f ` S \<subseteq> U\<close> \<open>retraction T S r\<close> in \<open>auto simp: continuous_on_compose2 retraction\<close>) |
|
901 |
qed |
|
902 |
||
903 |
lemma idempotent_imp_retraction: |
|
904 |
assumes "continuous_on S f" and "f ` S \<subseteq> S" and "\<And>x. x \<in> S \<Longrightarrow> f(f x) = f x" |
|
905 |
shows "retraction S (f ` S) f" |
|
906 |
by (simp add: assms retraction) |
|
907 |
||
908 |
lemma retraction_subset: |
|
909 |
assumes "retraction S T r" and "T \<subseteq> s'" and "s' \<subseteq> S" |
|
910 |
shows "retraction s' T r" |
|
911 |
unfolding retraction_def |
|
912 |
by (metis assms continuous_on_subset image_mono retraction) |
|
913 |
||
914 |
lemma retract_of_subset: |
|
915 |
assumes "T retract_of S" and "T \<subseteq> s'" and "s' \<subseteq> S" |
|
916 |
shows "T retract_of s'" |
|
917 |
by (meson assms retract_of_def retraction_subset) |
|
918 |
||
919 |
lemma retraction_refl [simp]: "retraction S S (\<lambda>x. x)" |
|
920 |
by (simp add: retraction) |
|
921 |
||
922 |
lemma retract_of_refl [iff]: "S retract_of S" |
|
923 |
unfolding retract_of_def retraction_def |
|
924 |
using continuous_on_id by blast |
|
925 |
||
926 |
lemma retract_of_imp_subset: |
|
927 |
"S retract_of T \<Longrightarrow> S \<subseteq> T" |
|
928 |
by (simp add: retract_of_def retraction_def) |
|
929 |
||
930 |
lemma retract_of_empty [simp]: |
|
931 |
"({} retract_of S) \<longleftrightarrow> S = {}" "(S retract_of {}) \<longleftrightarrow> S = {}"
|
|
932 |
by (auto simp: retract_of_def retraction_def) |
|
933 |
||
934 |
lemma retract_of_singleton [iff]: "({x} retract_of S) \<longleftrightarrow> x \<in> S"
|
|
935 |
unfolding retract_of_def retraction_def by force |
|
936 |
||
937 |
lemma retraction_comp: |
|
938 |
"\<lbrakk>retraction S T f; retraction T U g\<rbrakk> |
|
939 |
\<Longrightarrow> retraction S U (g \<circ> f)" |
|
940 |
apply (auto simp: retraction_def intro: continuous_on_compose2) |
|
941 |
by blast |
|
942 |
||
943 |
lemma retract_of_trans [trans]: |
|
944 |
assumes "S retract_of T" and "T retract_of U" |
|
945 |
shows "S retract_of U" |
|
946 |
using assms by (auto simp: retract_of_def intro: retraction_comp) |
|
947 |
||
948 |
lemma closedin_retract: |
|
949 |
fixes S :: "'a :: t2_space set" |
|
950 |
assumes "S retract_of T" |
|
951 |
shows "closedin (subtopology euclidean T) S" |
|
952 |
proof - |
|
953 |
obtain r where r: "S \<subseteq> T" "continuous_on T r" "r ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> r x = x" |
|
954 |
using assms by (auto simp: retract_of_def retraction_def) |
|
955 |
have "S = {x\<in>T. x = r x}"
|
|
956 |
using r by auto |
|
957 |
also have "\<dots> = T \<inter> ((\<lambda>x. (x, r x)) -` ({y. \<exists>x. y = (x, x)}))"
|
|
958 |
unfolding vimage_def mem_Times_iff fst_conv snd_conv |
|
959 |
using r |
|
960 |
by auto |
|
961 |
also have "closedin (top_of_set T) \<dots>" |
|
962 |
by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) |
|
963 |
finally show ?thesis . |
|
964 |
qed |
|
965 |
||
966 |
lemma closedin_self [simp]: "closedin (subtopology euclidean S) S" |
|
967 |
by simp |
|
968 |
||
969 |
lemma retract_of_closed: |
|
970 |
fixes S :: "'a :: t2_space set" |
|
971 |
shows "\<lbrakk>closed T; S retract_of T\<rbrakk> \<Longrightarrow> closed S" |
|
972 |
by (metis closedin_retract closedin_closed_eq) |
|
973 |
||
974 |
lemma retract_of_compact: |
|
975 |
"\<lbrakk>compact T; S retract_of T\<rbrakk> \<Longrightarrow> compact S" |
|
976 |
by (metis compact_continuous_image retract_of_def retraction) |
|
977 |
||
978 |
lemma retract_of_connected: |
|
979 |
"\<lbrakk>connected T; S retract_of T\<rbrakk> \<Longrightarrow> connected S" |
|
980 |
by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) |
|
981 |
||
982 |
lemma retraction_imp_quotient_map: |
|
983 |
"openin (subtopology euclidean S) (S \<inter> r -` U) \<longleftrightarrow> openin (subtopology euclidean T) U" |
|
984 |
if retraction: "retraction S T r" and "U \<subseteq> T" |
|
985 |
using retraction apply (rule retractionE) |
|
986 |
apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) |
|
987 |
using \<open>U \<subseteq> T\<close> apply (auto elim: continuous_on_subset) |
|
988 |
done |
|
989 |
||
990 |
lemma retract_of_Times: |
|
991 |
"\<lbrakk>S retract_of s'; T retract_of t'\<rbrakk> \<Longrightarrow> (S \<times> T) retract_of (s' \<times> t')" |
|
992 |
apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) |
|
993 |
apply (rename_tac f g) |
|
994 |
apply (rule_tac x="\<lambda>z. ((f \<circ> fst) z, (g \<circ> snd) z)" in exI) |
|
995 |
apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ |
|
996 |
done |
|
997 |
||
|
69616
d18dc9c5c456
split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff
changeset
|
998 |
end |