author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 77200 | 8f2e6186408f |
permissions | -rw-r--r-- |
77200
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1 |
section \<open>Some Uncountable Sets\<close> |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3 |
theory Uncountable_Sets |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
4 |
imports Path_Connected Continuum_Not_Denumerable |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
5 |
begin |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
6 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
7 |
lemma uncountable_closed_segment: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
8 |
fixes a :: "'a::real_normed_vector" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
9 |
assumes "a \<noteq> b" shows "uncountable (closed_segment a b)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
unfolding path_image_linepath [symmetric] path_image_def |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
11 |
using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1] |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
countable_image_inj_on by auto |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
13 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
lemma uncountable_open_segment: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
15 |
fixes a :: "'a::real_normed_vector" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
16 |
assumes "a \<noteq> b" shows "uncountable (open_segment a b)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
17 |
by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
18 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
lemma uncountable_convex: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
fixes a :: "'a::real_normed_vector" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
21 |
assumes "convex S" "a \<in> S" "b \<in> S" "a \<noteq> b" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
shows "uncountable S" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
23 |
proof - |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
24 |
have "uncountable (closed_segment a b)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
by (simp add: uncountable_closed_segment assms) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
then show ?thesis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
27 |
by (meson assms convex_contains_segment countable_subset) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
29 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
30 |
lemma uncountable_ball: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
31 |
fixes a :: "'a::euclidean_space" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
32 |
assumes "r > 0" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
33 |
shows "uncountable (ball a r)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
34 |
proof - |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
35 |
have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)))" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
36 |
by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
37 |
moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)) \<subseteq> ball a r" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
38 |
using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
39 |
ultimately show ?thesis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
40 |
by (metis countable_subset) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
42 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
43 |
lemma ball_minus_countable_nonempty: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
44 |
assumes "countable (A :: 'a :: euclidean_space set)" "r > 0" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
45 |
shows "ball z r - A \<noteq> {}" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
46 |
proof |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
47 |
assume *: "ball z r - A = {}" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
have "uncountable (ball z r - A)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
49 |
by (intro uncountable_minus_countable assms uncountable_ball) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
thus False by (subst (asm) *) auto |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
51 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
52 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
53 |
lemma uncountable_cball: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
fixes a :: "'a::euclidean_space" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
55 |
assumes "r > 0" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
56 |
shows "uncountable (cball a r)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
57 |
using assms countable_subset uncountable_ball by auto |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
58 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
59 |
lemma pairwise_disjnt_countable: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
60 |
fixes \<N> :: "nat set set" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
61 |
assumes "pairwise disjnt \<N>" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
shows "countable \<N>" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
63 |
by (simp add: assms countable_disjoint_open_subsets open_discrete) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
64 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
65 |
lemma pairwise_disjnt_countable_Union: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
66 |
assumes "countable (\<Union>\<N>)" and pwd: "pairwise disjnt \<N>" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
67 |
shows "countable \<N>" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
68 |
proof - |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
69 |
obtain f :: "_ \<Rightarrow> nat" where f: "inj_on f (\<Union>\<N>)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
70 |
using assms by blast |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
then have "pairwise disjnt (\<Union> X \<in> \<N>. {f ` X})" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
72 |
using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f]) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
73 |
then have "countable (\<Union> X \<in> \<N>. {f ` X})" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
74 |
using pairwise_disjnt_countable by blast |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
75 |
then show ?thesis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
76 |
by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
77 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
lemma connected_uncountable: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
fixes S :: "'a::metric_space set" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
assumes "connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
82 |
proof - |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
83 |
have "continuous_on S (dist a)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
84 |
by (intro continuous_intros) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
85 |
then have "connected (dist a ` S)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
86 |
by (metis connected_continuous_image \<open>connected S\<close>) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
87 |
then have "closed_segment 0 (dist a b) \<subseteq> (dist a ` S)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
88 |
by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
then have "uncountable (dist a ` S)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
90 |
by (metis \<open>a \<noteq> b\<close> countable_subset dist_eq_0_iff uncountable_closed_segment) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
91 |
then show ?thesis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
92 |
by blast |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
93 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
94 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
95 |
lemma path_connected_uncountable: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
96 |
fixes S :: "'a::metric_space set" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
97 |
assumes "path_connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
98 |
using path_connected_imp_connected assms connected_uncountable by metis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
99 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
100 |
lemma simple_path_image_uncountable: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
101 |
fixes g :: "real \<Rightarrow> 'a::metric_space" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
102 |
assumes "simple_path g" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
103 |
shows "uncountable (path_image g)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
104 |
proof - |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
have "g 0 \<in> path_image g" "g (1/2) \<in> path_image g" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
106 |
by (simp_all add: path_defs) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
moreover have "g 0 \<noteq> g (1/2)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
108 |
using assms by (fastforce simp add: simple_path_def loop_free_def) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
109 |
ultimately have "\<forall>a. \<not> path_image g \<subseteq> {a}" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
110 |
by blast |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
111 |
then show ?thesis |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
112 |
using assms connected_simple_path_image connected_uncountable by blast |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
113 |
qed |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
114 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
115 |
lemma arc_image_uncountable: |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
116 |
fixes g :: "real \<Rightarrow> 'a::metric_space" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
117 |
assumes "arc g" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
118 |
shows "uncountable (path_image g)" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
119 |
by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
120 |
|
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
121 |
end |