author | wenzelm |
Tue, 09 Jul 2024 12:32:33 +0200 | |
changeset 80548 | d1662f1296db |
parent 78336 | 6bae28577994 |
permissions | -rw-r--r-- |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1 |
section\<open>The binary product topology\<close> |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
2 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3 |
theory Product_Topology |
77939
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
77935
diff
changeset
|
4 |
imports Function_Topology |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
5 |
begin |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
6 |
|
71200
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents:
70178
diff
changeset
|
7 |
section \<open>Product Topology\<close> |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
8 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
9 |
subsection\<open>Definition\<close> |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
11 |
definition prod_topology :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<times> 'b) topology" where |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
"prod_topology X Y \<equiv> topology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
13 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
lemma open_product_open: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
15 |
assumes "open A" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
16 |
shows "\<exists>\<U>. \<U> \<subseteq> {S \<times> T |S T. open S \<and> open T} \<and> \<Union> \<U> = A" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
17 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
18 |
obtain f g where *: "\<And>u. u \<in> A \<Longrightarrow> open (f u) \<and> open (g u) \<and> u \<in> (f u) \<times> (g u) \<and> (f u) \<times> (g u) \<subseteq> A" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
using open_prod_def [of A] assms by metis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
let ?\<U> = "(\<lambda>u. f u \<times> g u) ` A" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
21 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
by (rule_tac x="?\<U>" in exI) (auto simp: dest: *) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
23 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
24 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
lemma open_product_open_eq: "(arbitrary union_of (\<lambda>U. \<exists>S T. U = S \<times> T \<and> open S \<and> open T)) = open" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
27 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
lemma openin_prod_topology: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
29 |
"openin (prod_topology X Y) = arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T})" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
30 |
unfolding prod_topology_def |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
31 |
proof (rule topology_inverse') |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
32 |
show "istopology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
33 |
apply (rule istopology_base, simp) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
34 |
by (metis openin_Int Times_Int_Times) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
35 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
36 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
37 |
lemma topspace_prod_topology [simp]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
38 |
"topspace (prod_topology X Y) = topspace X \<times> topspace Y" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
39 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
40 |
have "topspace(prod_topology X Y) = \<Union> (Collect (openin (prod_topology X Y)))" (is "_ = ?Z") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
unfolding topspace_def .. |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
42 |
also have "\<dots> = topspace X \<times> topspace Y" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
43 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
44 |
show "?Z \<subseteq> topspace X \<times> topspace Y" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
45 |
apply (auto simp: openin_prod_topology union_of_def arbitrary_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
46 |
using openin_subset by force+ |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
47 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
have *: "\<exists>A B. topspace X \<times> topspace Y = A \<times> B \<and> openin X A \<and> openin Y B" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
49 |
by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
show "topspace X \<times> topspace Y \<subseteq> ?Z" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
51 |
apply (rule Union_upper) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
52 |
using * by (simp add: openin_prod_topology arbitrary_union_of_inc) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
53 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
finally show ?thesis . |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
55 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
56 |
|
78336 | 57 |
lemma prod_topology_trivial_iff [simp]: |
58 |
"prod_topology X Y = trivial_topology \<longleftrightarrow> X = trivial_topology \<or> Y = trivial_topology" |
|
59 |
by (metis (full_types) Sigma_empty1 null_topspace_iff_trivial subset_empty times_subset_iff topspace_prod_topology) |
|
60 |
||
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
61 |
lemma subtopology_Times: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
shows "subtopology (prod_topology X Y) (S \<times> T) = prod_topology (subtopology X S) (subtopology Y T)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
63 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
64 |
have "((\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T) relative_to S \<times> T) = |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
65 |
(\<lambda>U. \<exists>S' T'. U = S' \<times> T' \<and> (openin X relative_to S) S' \<and> (openin Y relative_to T) T')" |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
66 |
by (auto simp: relative_to_def Times_Int_Times fun_eq_iff) metis |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
67 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
68 |
by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
69 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
70 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
lemma prod_topology_subtopology: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
72 |
"prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \<times> topspace Y)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
73 |
"prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \<times> T)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
74 |
by (auto simp: subtopology_Times) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
75 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
76 |
lemma prod_topology_discrete_topology: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
77 |
"discrete_topology (S \<times> T) = prod_topology (discrete_topology S) (discrete_topology T)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
by (simp add: prod_topology_def open_product_open_eq) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
82 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
83 |
lemma prod_topology_subtopology_eu [simp]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
84 |
"prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \<times> T)" |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
85 |
by (simp add: prod_topology_subtopology subtopology_subtopology Times_Int_Times) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
86 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
87 |
lemma openin_prod_topology_alt: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
88 |
"openin (prod_topology X Y) S \<longleftrightarrow> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
(\<forall>x y. (x,y) \<in> S \<longrightarrow> (\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> S))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
90 |
apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
91 |
by (metis mem_Sigma_iff) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
92 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
93 |
lemma open_map_fst: "open_map (prod_topology X Y) X fst" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
94 |
unfolding open_map_def openin_prod_topology_alt |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
95 |
by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
96 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
97 |
lemma open_map_snd: "open_map (prod_topology X Y) Y snd" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
98 |
unfolding open_map_def openin_prod_topology_alt |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
99 |
by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
100 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
101 |
lemma openin_prod_Times_iff: |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
102 |
"openin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> openin X S \<and> openin Y T" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
103 |
proof (cases "S = {} \<or> T = {}") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
104 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
106 |
apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
apply (meson|force)+ |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
108 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
109 |
qed force |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
110 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
111 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
112 |
lemma closure_of_Times: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
113 |
"(prod_topology X Y) closure_of (S \<times> T) = (X closure_of S) \<times> (Y closure_of T)" (is "?lhs = ?rhs") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
114 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
115 |
show "?lhs \<subseteq> ?rhs" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
116 |
by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
117 |
show "?rhs \<subseteq> ?lhs" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
118 |
by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
119 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
120 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
121 |
lemma closedin_prod_Times_iff: |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
122 |
"closedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> closedin X S \<and> closedin Y T" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
123 |
by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
124 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
125 |
lemma interior_of_Times: "(prod_topology X Y) interior_of (S \<times> T) = (X interior_of S) \<times> (Y interior_of T)" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
126 |
proof (rule interior_of_unique) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
127 |
show "(X interior_of S) \<times> Y interior_of T \<subseteq> S \<times> T" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
128 |
by (simp add: Sigma_mono interior_of_subset) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
129 |
show "openin (prod_topology X Y) ((X interior_of S) \<times> Y interior_of T)" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
130 |
by (simp add: openin_prod_Times_iff) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
131 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
132 |
show "T' \<subseteq> (X interior_of S) \<times> Y interior_of T" if "T' \<subseteq> S \<times> T" "openin (prod_topology X Y) T'" for T' |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
133 |
proof (clarsimp; intro conjI) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
134 |
fix a :: "'a" and b :: "'b" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
135 |
assume "(a, b) \<in> T'" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
136 |
with that obtain U V where UV: "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> T'" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
137 |
by (metis openin_prod_topology_alt) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
138 |
then show "a \<in> X interior_of S" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
139 |
using interior_of_maximal_eq that(1) by fastforce |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
140 |
show "b \<in> Y interior_of T" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
141 |
using UV interior_of_maximal_eq that(1) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
142 |
by (metis SigmaI mem_Sigma_iff subset_eq) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
143 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
144 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
145 |
|
78093
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
146 |
text \<open>Missing the opposite direction. Does it hold? A converse is proved for proper maps, a stronger condition\<close> |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
147 |
lemma closed_map_prod: |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
148 |
assumes "closed_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y))" |
78336 | 149 |
shows "(prod_topology X Y) = trivial_topology \<or> closed_map X X' f \<and> closed_map Y Y' g" |
150 |
proof (cases "(prod_topology X Y) = trivial_topology") |
|
78093
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
151 |
case False |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
152 |
then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}" |
78336 | 153 |
by (auto simp flip: null_topspace_iff_trivial) |
78093
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
154 |
have "closed_map X X' f" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
155 |
unfolding closed_map_def |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
156 |
proof (intro strip) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
157 |
fix C |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
158 |
assume "closedin X C" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
159 |
show "closedin X' (f ` C)" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
160 |
proof (cases "C={}") |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
161 |
case False |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
162 |
with assms have "closedin (prod_topology X' Y') ((\<lambda>(x,y). (f x, g y)) ` (C \<times> topspace Y))" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
163 |
by (simp add: \<open>closedin X C\<close> closed_map_def closedin_prod_Times_iff) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
164 |
with False ne show ?thesis |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
165 |
by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
166 |
qed auto |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
167 |
qed |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
168 |
moreover |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
169 |
have "closed_map Y Y' g" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
170 |
unfolding closed_map_def |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
171 |
proof (intro strip) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
172 |
fix C |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
173 |
assume "closedin Y C" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
174 |
show "closedin Y' (g ` C)" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
175 |
proof (cases "C={}") |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
176 |
case False |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
177 |
with assms have "closedin (prod_topology X' Y') ((\<lambda>(x,y). (f x, g y)) ` (topspace X \<times> C))" |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
178 |
by (simp add: \<open>closedin Y C\<close> closed_map_def closedin_prod_Times_iff) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
179 |
with False ne show ?thesis |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
180 |
by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
181 |
qed auto |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
182 |
qed |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
183 |
ultimately show ?thesis |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
184 |
by (auto simp: False) |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
185 |
qed auto |
cec875dcc59e
Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents:
78037
diff
changeset
|
186 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
187 |
subsection \<open>Continuity\<close> |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
188 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
189 |
lemma continuous_map_pairwise: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
190 |
"continuous_map Z (prod_topology X Y) f \<longleftrightarrow> continuous_map Z X (fst \<circ> f) \<and> continuous_map Z Y (snd \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
191 |
(is "?lhs = ?rhs") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
192 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
193 |
let ?g = "fst \<circ> f" and ?h = "snd \<circ> f" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
194 |
have f: "f x = (?g x, ?h x)" for x |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
195 |
by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
196 |
show ?thesis |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
197 |
proof (cases "?g \<in> topspace Z \<rightarrow> topspace X \<and> ?h \<in> topspace Z \<rightarrow> topspace Y") |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
198 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
199 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
200 |
proof safe |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
201 |
assume "continuous_map Z (prod_topology X Y) f" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
202 |
then have "openin Z {x \<in> topspace Z. fst (f x) \<in> U}" if "openin X U" for U |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
203 |
unfolding continuous_map_def using True that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
204 |
apply clarify |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
205 |
apply (drule_tac x="U \<times> topspace Y" in spec) |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
206 |
by (auto simp: openin_prod_Times_iff mem_Times_iff Pi_iff cong: conj_cong) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
with True show "continuous_map Z X (fst \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
208 |
by (auto simp: continuous_map_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
209 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
210 |
assume "continuous_map Z (prod_topology X Y) f" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
211 |
then have "openin Z {x \<in> topspace Z. snd (f x) \<in> V}" if "openin Y V" for V |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
212 |
unfolding continuous_map_def using True that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
213 |
apply clarify |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
214 |
apply (drule_tac x="topspace X \<times> V" in spec) |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
215 |
by (simp add: openin_prod_Times_iff mem_Times_iff Pi_iff cong: conj_cong) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
216 |
with True show "continuous_map Z Y (snd \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
217 |
by (auto simp: continuous_map_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
218 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
219 |
assume Z: "continuous_map Z X (fst \<circ> f)" "continuous_map Z Y (snd \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
220 |
have *: "openin Z {x \<in> topspace Z. f x \<in> W}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
221 |
if "\<And>w. w \<in> W \<Longrightarrow> \<exists>U V. openin X U \<and> openin Y V \<and> w \<in> U \<times> V \<and> U \<times> V \<subseteq> W" for W |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
proof (subst openin_subopen, clarify) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
223 |
fix x :: "'a" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
224 |
assume "x \<in> topspace Z" and "f x \<in> W" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
225 |
with that [OF \<open>f x \<in> W\<close>] |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
226 |
obtain U V where UV: "openin X U" "openin Y V" "f x \<in> U \<times> V" "U \<times> V \<subseteq> W" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
227 |
by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
228 |
with Z UV show "\<exists>T. openin Z T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace Z. f x \<in> W}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
229 |
apply (rule_tac x="{x \<in> topspace Z. ?g x \<in> U} \<inter> {x \<in> topspace Z. ?h x \<in> V}" in exI) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
230 |
apply (auto simp: \<open>x \<in> topspace Z\<close> continuous_map_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
231 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
232 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
233 |
show "continuous_map Z (prod_topology X Y) f" |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
234 |
using True by (force simp: continuous_map_def openin_prod_topology_alt mem_Times_iff *) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
235 |
qed |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
236 |
qed (force simp: continuous_map_def) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
237 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
238 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
239 |
lemma continuous_map_paired: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
240 |
"continuous_map Z (prod_topology X Y) (\<lambda>x. (f x,g x)) \<longleftrightarrow> continuous_map Z X f \<and> continuous_map Z Y g" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
241 |
by (simp add: continuous_map_pairwise o_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
242 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
243 |
lemma continuous_map_pairedI [continuous_intros]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
244 |
"\<lbrakk>continuous_map Z X f; continuous_map Z Y g\<rbrakk> \<Longrightarrow> continuous_map Z (prod_topology X Y) (\<lambda>x. (f x,g x))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
245 |
by (simp add: continuous_map_pairwise o_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
246 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
247 |
lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
248 |
using continuous_map_pairwise [of "prod_topology X Y" X Y id] |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
249 |
by (simp add: continuous_map_pairwise) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
251 |
lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
252 |
using continuous_map_pairwise [of "prod_topology X Y" X Y id] |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
253 |
by (simp add: continuous_map_pairwise) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
254 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
255 |
lemma continuous_map_fst_of [continuous_intros]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
256 |
"continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z X (fst \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
257 |
by (simp add: continuous_map_pairwise) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
258 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
259 |
lemma continuous_map_snd_of [continuous_intros]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
260 |
"continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z Y (snd \<circ> f)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
261 |
by (simp add: continuous_map_pairwise) |
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
262 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
263 |
lemma continuous_map_prod_fst: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
264 |
"i \<in> I \<Longrightarrow> continuous_map (prod_topology (product_topology (\<lambda>i. Y) I) X) Y (\<lambda>x. fst x i)" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
265 |
using continuous_map_componentwise_UNIV continuous_map_fst by fastforce |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
266 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
267 |
lemma continuous_map_prod_snd: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
268 |
"i \<in> I \<Longrightarrow> continuous_map (prod_topology X (product_topology (\<lambda>i. Y) I)) Y (\<lambda>x. snd x i)" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
269 |
using continuous_map_componentwise_UNIV continuous_map_snd by fastforce |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
270 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
271 |
lemma continuous_map_if_iff [simp]: "continuous_map X Y (\<lambda>x. if P then f x else g x) \<longleftrightarrow> continuous_map X Y (if P then f else g)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
272 |
by simp |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
273 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
274 |
lemma continuous_map_if [continuous_intros]: "\<lbrakk>P \<Longrightarrow> continuous_map X Y f; ~P \<Longrightarrow> continuous_map X Y g\<rbrakk> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
275 |
\<Longrightarrow> continuous_map X Y (\<lambda>x. if P then f x else g x)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
276 |
by simp |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
277 |
|
78336 | 278 |
lemma prod_topology_trivial1 [simp]: "prod_topology trivial_topology Y = trivial_topology" |
279 |
using continuous_map_fst continuous_map_on_empty2 by blast |
|
280 |
||
281 |
lemma prod_topology_trivial2 [simp]: "prod_topology X trivial_topology = trivial_topology" |
|
282 |
using continuous_map_snd continuous_map_on_empty2 by blast |
|
283 |
||
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
284 |
lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
285 |
using continuous_map_from_subtopology continuous_map_fst by force |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
286 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
287 |
lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
288 |
using continuous_map_from_subtopology continuous_map_snd by force |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
289 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
290 |
lemma quotient_map_fst [simp]: |
78336 | 291 |
"quotient_map(prod_topology X Y) X fst \<longleftrightarrow> (Y = trivial_topology \<longrightarrow> X = trivial_topology)" |
292 |
apply (simp add: continuous_open_quotient_map open_map_fst continuous_map_fst) |
|
293 |
by (metis null_topspace_iff_trivial) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
294 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
295 |
lemma quotient_map_snd [simp]: |
78336 | 296 |
"quotient_map(prod_topology X Y) Y snd \<longleftrightarrow> (X = trivial_topology \<longrightarrow> Y = trivial_topology)" |
297 |
apply (simp add: continuous_open_quotient_map open_map_snd continuous_map_snd) |
|
298 |
by (metis null_topspace_iff_trivial) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
299 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
300 |
lemma retraction_map_fst: |
78336 | 301 |
"retraction_map (prod_topology X Y) X fst \<longleftrightarrow> (Y = trivial_topology \<longrightarrow> X = trivial_topology)" |
302 |
proof (cases "Y = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
303 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
304 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
305 |
using continuous_map_image_subset_topspace |
78336 | 306 |
by (auto simp: retraction_map_def retraction_maps_def continuous_map_pairwise) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
307 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
308 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
309 |
have "\<exists>g. continuous_map X (prod_topology X Y) g \<and> (\<forall>x\<in>topspace X. fst (g x) = x)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
310 |
if y: "y \<in> topspace Y" for y |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
311 |
by (rule_tac x="\<lambda>x. (x,y)" in exI) (auto simp: y continuous_map_paired) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
312 |
with False have "retraction_map (prod_topology X Y) X fst" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
314 |
with False show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
315 |
by simp |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
316 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
317 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
318 |
lemma retraction_map_snd: |
78336 | 319 |
"retraction_map (prod_topology X Y) Y snd \<longleftrightarrow> (X = trivial_topology \<longrightarrow> Y = trivial_topology)" |
320 |
proof (cases "X = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
321 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
322 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
323 |
using continuous_map_image_subset_topspace |
78336 | 324 |
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
325 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
326 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
327 |
have "\<exists>g. continuous_map Y (prod_topology X Y) g \<and> (\<forall>y\<in>topspace Y. snd (g y) = y)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
328 |
if x: "x \<in> topspace X" for x |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
329 |
by (rule_tac x="\<lambda>y. (x,y)" in exI) (auto simp: x continuous_map_paired) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
330 |
with False have "retraction_map (prod_topology X Y) Y snd" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
331 |
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
332 |
with False show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
333 |
by simp |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
334 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
335 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
336 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
337 |
lemma continuous_map_of_fst: |
78336 | 338 |
"continuous_map (prod_topology X Y) Z (f \<circ> fst) \<longleftrightarrow> Y = trivial_topology \<or> continuous_map X Z f" |
339 |
proof (cases "Y = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
340 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
341 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
342 |
by (simp add: continuous_map_on_empty) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
344 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
345 |
then show ?thesis |
71633 | 346 |
by (simp add: continuous_compose_quotient_map_eq) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
347 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
348 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
349 |
lemma continuous_map_of_snd: |
78336 | 350 |
"continuous_map (prod_topology X Y) Z (f \<circ> snd) \<longleftrightarrow> X = trivial_topology \<or> continuous_map Y Z f" |
351 |
proof (cases "X = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
352 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
353 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
354 |
by (simp add: continuous_map_on_empty) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
355 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
356 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
357 |
then show ?thesis |
71633 | 358 |
by (simp add: continuous_compose_quotient_map_eq) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
359 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
360 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
361 |
lemma continuous_map_prod_top: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
362 |
"continuous_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow> |
78336 | 363 |
(prod_topology X Y) = trivial_topology \<or> continuous_map X X' f \<and> continuous_map Y Y' g" |
364 |
proof (cases "(prod_topology X Y) = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
365 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
366 |
then show ?thesis |
78336 | 367 |
by (auto simp: continuous_map_paired case_prod_unfold |
368 |
continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def]) |
|
369 |
qed auto |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
370 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
371 |
lemma in_prod_topology_closure_of: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
372 |
assumes "z \<in> (prod_topology X Y) closure_of S" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
373 |
shows "fst z \<in> X closure_of (fst ` S)" "snd z \<in> Y closure_of (snd ` S)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
374 |
using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
375 |
using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
376 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
377 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
378 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
379 |
proposition compact_space_prod_topology: |
78336 | 380 |
"compact_space(prod_topology X Y) \<longleftrightarrow> (prod_topology X Y) = trivial_topology \<or> compact_space X \<and> compact_space Y" |
381 |
proof (cases "(prod_topology X Y) = trivial_topology") |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
382 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
383 |
then show ?thesis |
78336 | 384 |
by fastforce |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
385 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
386 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
387 |
then have non_mt: "topspace X \<noteq> {}" "topspace Y \<noteq> {}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
388 |
by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
389 |
have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
390 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
391 |
have "compactin X (fst ` (topspace X \<times> topspace Y))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
392 |
by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
393 |
moreover |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
394 |
have "compactin Y (snd ` (topspace X \<times> topspace Y))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
395 |
by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
ultimately show "compact_space X" "compact_space Y" |
78336 | 397 |
using non_mt by (auto simp: compact_space_def) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
398 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
399 |
moreover |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
400 |
define \<X> where "\<X> \<equiv> (\<lambda>V. topspace X \<times> V) ` Collect (openin Y)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
401 |
define \<Y> where "\<Y> \<equiv> (\<lambda>U. U \<times> topspace Y) ` Collect (openin X)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
402 |
have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
403 |
proof (rule Alexander_subbase_alt) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
404 |
show toptop: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X> \<union> \<Y>)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
405 |
unfolding \<X>_def \<Y>_def by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
406 |
fix \<C> :: "('a \<times> 'b) set set" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
407 |
assume \<C>: "\<C> \<subseteq> \<X> \<union> \<Y>" "topspace X \<times> topspace Y \<subseteq> \<Union>\<C>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
408 |
then obtain \<X>' \<Y>' where XY: "\<X>' \<subseteq> \<X>" "\<Y>' \<subseteq> \<Y>" and \<C>eq: "\<C> = \<X>' \<union> \<Y>'" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
409 |
using subset_UnE by metis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
410 |
then have sub: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X>' \<union> \<Y>')" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
411 |
using \<C> by simp |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
412 |
obtain \<U> \<V> where \<U>: "\<And>U. U \<in> \<U> \<Longrightarrow> openin X U" "\<Y>' = (\<lambda>U. U \<times> topspace Y) ` \<U>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
413 |
and \<V>: "\<And>V. V \<in> \<V> \<Longrightarrow> openin Y V" "\<X>' = (\<lambda>V. topspace X \<times> V) ` \<V>" |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
414 |
using XY by (clarsimp simp add: \<X>_def \<Y>_def subset_image_iff) (force simp: subset_iff) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
415 |
have "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<X>' \<union> \<Y>' \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
416 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
417 |
have "topspace X \<subseteq> \<Union>\<U> \<or> topspace Y \<subseteq> \<Union>\<V>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
418 |
using \<U> \<V> \<C> \<C>eq by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
419 |
then have *: "\<exists>\<D>. finite \<D> \<and> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
420 |
(\<forall>x \<in> \<D>. x \<in> (\<lambda>V. topspace X \<times> V) ` \<V> \<or> x \<in> (\<lambda>U. U \<times> topspace Y) ` \<U>) \<and> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
421 |
(topspace X \<times> topspace Y \<subseteq> \<Union>\<D>)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
422 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
423 |
assume "topspace X \<subseteq> \<Union>\<U>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
424 |
with \<open>compact_space X\<close> \<U> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<U>" "topspace X \<subseteq> \<Union>\<F>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
425 |
by (meson compact_space_alt) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
426 |
with that show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
427 |
by (rule_tac x="(\<lambda>D. D \<times> topspace Y) ` \<F>" in exI) auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
428 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
429 |
assume "topspace Y \<subseteq> \<Union>\<V>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
430 |
with \<open>compact_space Y\<close> \<V> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "topspace Y \<subseteq> \<Union>\<F>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
431 |
by (meson compact_space_alt) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
432 |
with that show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
433 |
by (rule_tac x="(\<lambda>C. topspace X \<times> C) ` \<F>" in exI) auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
434 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
435 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
436 |
using that \<U> \<V> by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
437 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
438 |
then show "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<C> \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
439 |
using \<C> \<C>eq by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
440 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
441 |
have "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>) relative_to topspace X \<times> topspace Y) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
442 |
= (\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
443 |
(is "?lhs = ?rhs") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
444 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
445 |
have "?rhs U" if "?lhs U" for U |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
446 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
447 |
have "topspace X \<times> topspace Y \<inter> \<Inter> T \<in> {A \<times> B |A B. A \<in> Collect (openin X) \<and> B \<in> Collect (openin Y)}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
448 |
if "finite T" "T \<subseteq> \<X> \<union> \<Y>" for T |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
449 |
using that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
450 |
proof induction |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
451 |
case (insert B \<B>) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
452 |
then show ?case |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
453 |
unfolding \<X>_def \<Y>_def |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
454 |
apply (simp add: Int_ac subset_eq image_def) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
455 |
apply (metis (no_types) openin_Int openin_topspace Times_Int_Times) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
456 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
457 |
qed auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
458 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
459 |
using that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
460 |
by (auto simp: subset_eq elim!: relative_toE intersection_ofE) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
461 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
462 |
moreover |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
463 |
have "?lhs Z" if Z: "?rhs Z" for Z |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
464 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
465 |
obtain U V where "Z = U \<times> V" "openin X U" "openin Y V" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
466 |
using Z by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
467 |
then have UV: "U \<times> V = (topspace X \<times> topspace Y) \<inter> (U \<times> V)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
468 |
by (simp add: Sigma_mono inf_absorb2 openin_subset) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
469 |
moreover |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
470 |
have "?lhs ((topspace X \<times> topspace Y) \<inter> (U \<times> V))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
471 |
proof (rule relative_to_inc) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
472 |
show "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>)) (U \<times> V)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
473 |
apply (simp add: intersection_of_def \<X>_def \<Y>_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
474 |
apply (rule_tac x="{(U \<times> topspace Y),(topspace X \<times> V)}" in exI) |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
475 |
using \<open>openin X U\<close> \<open>openin Y V\<close> openin_subset UV apply (fastforce simp:) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
476 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
477 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
478 |
ultimately show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
479 |
using \<open>Z = U \<times> V\<close> by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
480 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
481 |
ultimately show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
482 |
by meson |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
483 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
484 |
then show "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<X> \<union> \<Y>) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
485 |
relative_to (topspace X \<times> topspace Y))) = |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
486 |
prod_topology X Y" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
487 |
by (simp add: prod_topology_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
488 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
489 |
ultimately show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
490 |
using False by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
491 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
492 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
493 |
lemma compactin_Times: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
494 |
"compactin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> compactin X S \<and> compactin Y T" |
78336 | 495 |
by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology subtopology_trivial_iff) |
496 |
||
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
497 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
498 |
subsection\<open>Homeomorphic maps\<close> |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
499 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
500 |
lemma homeomorphic_maps_prod: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
501 |
"homeomorphic_maps (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) (\<lambda>(x,y). (f' x, g' y)) \<longleftrightarrow> |
78336 | 502 |
(prod_topology X Y) = trivial_topology \<and> (prod_topology X' Y') = trivial_topology |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
503 |
\<or> homeomorphic_maps X X' f f' \<and> homeomorphic_maps Y Y' g g'" (is "?lhs = ?rhs") |
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
504 |
proof |
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
505 |
show "?lhs \<Longrightarrow> ?rhs" |
78336 | 506 |
by (fastforce simp: homeomorphic_maps_def continuous_map_prod_top ball_conj_distrib) |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
507 |
next |
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
508 |
show "?rhs \<Longrightarrow> ?lhs" |
78336 | 509 |
by (auto simp: homeomorphic_maps_def continuous_map_prod_top) |
78320
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
510 |
qed |
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents:
78093
diff
changeset
|
511 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
512 |
|
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
513 |
lemma homeomorphic_maps_swap: |
78336 | 514 |
"homeomorphic_maps (prod_topology X Y) (prod_topology Y X) (\<lambda>(x,y). (y,x)) (\<lambda>(y,x). (x,y))" |
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
515 |
by (auto simp: homeomorphic_maps_def case_prod_unfold continuous_map_fst continuous_map_pairedI continuous_map_snd) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
516 |
|
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
517 |
lemma homeomorphic_map_swap: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
518 |
"homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(x,y). (y,x))" |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
519 |
using homeomorphic_map_maps homeomorphic_maps_swap by metis |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
520 |
|
77935
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
521 |
lemma homeomorphic_space_prod_topology_swap: |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
522 |
"(prod_topology X Y) homeomorphic_space (prod_topology Y X)" |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
523 |
using homeomorphic_map_swap homeomorphic_space by blast |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
524 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
525 |
lemma embedding_map_graph: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
526 |
"embedding_map X (prod_topology X Y) (\<lambda>x. (x, f x)) \<longleftrightarrow> continuous_map X Y f" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
527 |
(is "?lhs = ?rhs") |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
528 |
proof |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
529 |
assume L: ?lhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
530 |
have "snd \<circ> (\<lambda>x. (x, f x)) = f" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
531 |
by force |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
532 |
moreover have "continuous_map X Y (snd \<circ> (\<lambda>x. (x, f x)))" |
77935
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
533 |
using L unfolding embedding_map_def |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
534 |
by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
535 |
ultimately show ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
536 |
by simp |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
537 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
538 |
assume R: ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
539 |
then show ?lhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
540 |
unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
541 |
by (rule_tac x=fst in exI) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
542 |
(auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology |
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
543 |
continuous_map_fst) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
544 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
545 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
546 |
lemma homeomorphic_space_prod_topology: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
547 |
"\<lbrakk>X homeomorphic_space X''; Y homeomorphic_space Y'\<rbrakk> |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
548 |
\<Longrightarrow> prod_topology X Y homeomorphic_space prod_topology X'' Y'" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
549 |
using homeomorphic_maps_prod unfolding homeomorphic_space_def by blast |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
550 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
551 |
lemma prod_topology_homeomorphic_space_left: |
78336 | 552 |
"Y = discrete_topology {b} \<Longrightarrow> prod_topology X Y homeomorphic_space X" |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
553 |
unfolding homeomorphic_space_def |
78336 | 554 |
apply (rule_tac x=fst in exI) |
555 |
apply (simp add: homeomorphic_map_def inj_on_def discrete_topology_unique flip: homeomorphic_map_maps) |
|
556 |
done |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
557 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
558 |
lemma prod_topology_homeomorphic_space_right: |
78336 | 559 |
"X = discrete_topology {a} \<Longrightarrow> prod_topology X Y homeomorphic_space Y" |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
560 |
unfolding homeomorphic_space_def |
78336 | 561 |
by (meson homeomorphic_space_def homeomorphic_space_prod_topology_swap homeomorphic_space_trans prod_topology_homeomorphic_space_left) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
562 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
563 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
564 |
lemma homeomorphic_space_prod_topology_sing1: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
565 |
"b \<in> topspace Y \<Longrightarrow> X homeomorphic_space (prod_topology X (subtopology Y {b}))" |
78336 | 566 |
by (metis empty_subsetI homeomorphic_space_sym insert_subset prod_topology_homeomorphic_space_left subtopology_eq_discrete_topology_sing topspace_subtopology_subset) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
567 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
568 |
lemma homeomorphic_space_prod_topology_sing2: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
569 |
"a \<in> topspace X \<Longrightarrow> Y homeomorphic_space (prod_topology (subtopology X {a}) Y)" |
78336 | 570 |
by (metis empty_subsetI homeomorphic_space_sym insert_subset prod_topology_homeomorphic_space_right subtopology_eq_discrete_topology_sing topspace_subtopology_subset) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
571 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
572 |
lemma topological_property_of_prod_component: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
573 |
assumes major: "P(prod_topology X Y)" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
574 |
and X: "\<And>x. \<lbrakk>x \<in> topspace X; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) ({x} \<times> topspace Y))" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
575 |
and Y: "\<And>y. \<lbrakk>y \<in> topspace Y; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) (topspace X \<times> {y}))" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
576 |
and PQ: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
577 |
and PR: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> R X')" |
78336 | 578 |
shows "(prod_topology X Y) = trivial_topology \<or> Q X \<and> R Y" |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
579 |
proof - |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
580 |
have "Q X \<and> R Y" if "topspace(prod_topology X Y) \<noteq> {}" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
581 |
proof - |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
582 |
from that obtain a b where a: "a \<in> topspace X" and b: "b \<in> topspace Y" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
583 |
by force |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
584 |
show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
585 |
using X [OF a major] and Y [OF b major] homeomorphic_space_prod_topology_sing1 [OF b, of X] homeomorphic_space_prod_topology_sing2 [OF a, of Y] |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
586 |
by (simp add: subtopology_Times) (meson PQ PR homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
587 |
qed |
78336 | 588 |
then show ?thesis by force |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
589 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
590 |
|
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
591 |
lemma limitin_pairwise: |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
592 |
"limitin (prod_topology X Y) f l F \<longleftrightarrow> limitin X (fst \<circ> f) (fst l) F \<and> limitin Y (snd \<circ> f) (snd l) F" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
593 |
(is "?lhs = ?rhs") |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
594 |
proof |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
595 |
assume ?lhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
596 |
then obtain a b where ev: "\<And>U. \<lbrakk>(a,b) \<in> U; openin (prod_topology X Y) U\<rbrakk> \<Longrightarrow> \<forall>\<^sub>F x in F. f x \<in> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
597 |
and a: "a \<in> topspace X" and b: "b \<in> topspace Y" and l: "l = (a,b)" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
598 |
by (auto simp: limitin_def) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
599 |
moreover have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" if "openin X U" "a \<in> U" for U |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
600 |
proof - |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
601 |
have "\<forall>\<^sub>F c in F. f c \<in> U \<times> topspace Y" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
602 |
using b that ev [of "U \<times> topspace Y"] by (auto simp: openin_prod_topology_alt) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
603 |
then show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
604 |
by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
605 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
606 |
moreover have "\<forall>\<^sub>F x in F. snd (f x) \<in> U" if "openin Y U" "b \<in> U" for U |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
607 |
proof - |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
608 |
have "\<forall>\<^sub>F c in F. f c \<in> topspace X \<times> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
609 |
using a that ev [of "topspace X \<times> U"] by (auto simp: openin_prod_topology_alt) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
610 |
then show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
611 |
by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
612 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
613 |
ultimately show ?rhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
614 |
by (simp add: limitin_def) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
615 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
616 |
have "limitin (prod_topology X Y) f (a,b) F" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
617 |
if "limitin X (fst \<circ> f) a F" "limitin Y (snd \<circ> f) b F" for a b |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
618 |
using that |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
619 |
proof (clarsimp simp: limitin_def) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
620 |
fix Z :: "('a \<times> 'b) set" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
621 |
assume a: "a \<in> topspace X" "\<forall>U. openin X U \<and> a \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. fst (f x) \<in> U)" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
622 |
and b: "b \<in> topspace Y" "\<forall>U. openin Y U \<and> b \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. snd (f x) \<in> U)" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
623 |
and Z: "openin (prod_topology X Y) Z" "(a, b) \<in> Z" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
624 |
then obtain U V where "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> Z" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
625 |
using Z by (force simp: openin_prod_topology_alt) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
626 |
then have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" "\<forall>\<^sub>F x in F. snd (f x) \<in> V" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
627 |
by (simp_all add: a b) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
628 |
then show "\<forall>\<^sub>F x in F. f x \<in> Z" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
629 |
by (rule eventually_elim2) (use \<open>U \<times> V \<subseteq> Z\<close> subsetD in auto) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
630 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
631 |
then show "?rhs \<Longrightarrow> ?lhs" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
632 |
by (metis prod.collapse) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
633 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
634 |
|
78037
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
635 |
proposition connected_space_prod_topology: |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
636 |
"connected_space(prod_topology X Y) \<longleftrightarrow> |
78336 | 637 |
(prod_topology X Y) = trivial_topology \<or> connected_space X \<and> connected_space Y" (is "?lhs=?rhs") |
638 |
proof (cases "(prod_topology X Y) = trivial_topology") |
|
78037
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
639 |
case True |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
640 |
then show ?thesis |
78336 | 641 |
by auto |
78037
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
642 |
next |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
643 |
case False |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
644 |
then have nonempty: "topspace X \<noteq> {}" "topspace Y \<noteq> {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
645 |
by force+ |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
646 |
show ?thesis |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
647 |
proof |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
648 |
assume ?lhs |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
649 |
then show ?rhs |
78336 | 650 |
by (metis connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd |
651 |
subtopology_eq_discrete_topology_empty) |
|
78037
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
652 |
next |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
653 |
assume ?rhs |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
654 |
then have conX: "connected_space X" and conY: "connected_space Y" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
655 |
using False by blast+ |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
656 |
have False |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
657 |
if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
658 |
and UV: "topspace X \<times> topspace Y \<subseteq> U \<union> V" "U \<inter> V = {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
659 |
and "U \<noteq> {}" and "V \<noteq> {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
660 |
for U V |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
661 |
proof - |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
662 |
have Usub: "U \<subseteq> topspace X \<times> topspace Y" and Vsub: "V \<subseteq> topspace X \<times> topspace Y" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
663 |
using that by (metis openin_subset topspace_prod_topology)+ |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
664 |
obtain a b where ab: "(a,b) \<in> U" and a: "a \<in> topspace X" and b: "b \<in> topspace Y" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
665 |
using \<open>U \<noteq> {}\<close> Usub by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
666 |
have "\<not> topspace X \<times> topspace Y \<subseteq> U" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
667 |
using Usub Vsub \<open>U \<inter> V = {}\<close> \<open>V \<noteq> {}\<close> by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
668 |
then obtain x y where x: "x \<in> topspace X" and y: "y \<in> topspace Y" and "(x,y) \<notin> U" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
669 |
by blast |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
670 |
have oX: "openin X {x \<in> topspace X. (x,y) \<in> U}" "openin X {x \<in> topspace X. (x,y) \<in> V}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
671 |
and oY: "openin Y {y \<in> topspace Y. (a,y) \<in> U}" "openin Y {y \<in> topspace Y. (a,y) \<in> V}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
672 |
by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
673 |
simp: that continuous_map_pairwise o_def x y a)+ |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
674 |
have 1: "topspace Y \<subseteq> {y \<in> topspace Y. (a,y) \<in> U} \<union> {y \<in> topspace Y. (a,y) \<in> V}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
675 |
using a that(3) by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
676 |
have 2: "{y \<in> topspace Y. (a,y) \<in> U} \<inter> {y \<in> topspace Y. (a,y) \<in> V} = {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
677 |
using that(4) by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
678 |
have 3: "{y \<in> topspace Y. (a,y) \<in> U} \<noteq> {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
679 |
using ab b by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
680 |
have 4: "{y \<in> topspace Y. (a,y) \<in> V} \<noteq> {}" |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
681 |
proof - |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
682 |
show ?thesis |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
683 |
using connected_spaceD [OF conX oX] UV \<open>(x,y) \<notin> U\<close> a x y |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
684 |
disjoint_iff_not_equal by blast |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
685 |
qed |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
686 |
show ?thesis |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
687 |
using connected_spaceD [OF conY oY 1 2 3 4] by auto |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
688 |
qed |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
689 |
then show ?lhs |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
690 |
unfolding connected_space_def topspace_prod_topology by blast |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
691 |
qed |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
692 |
qed |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
693 |
|
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
694 |
lemma connectedin_Times: |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
695 |
"connectedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> |
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
696 |
S = {} \<or> T = {} \<or> connectedin X S \<and> connectedin Y T" |
78336 | 697 |
by (auto simp: connectedin_def subtopology_Times connected_space_prod_topology subtopology_trivial_iff) |
78037
37894dff0111
More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents:
77939
diff
changeset
|
698 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
699 |
end |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
700 |