| author | wenzelm | 
| Thu, 26 Mar 2020 11:48:52 +0100 | |
| changeset 71663 | fb7fdd3eb7b9 | 
| parent 70136 | f03a01a18c6e | 
| child 75078 | ec86cb2418e1 | 
| permissions | -rw-r--r-- | 
| 
63970
 
3b6a3632e754
HOL-Analysis: move Continuum_Not_Denumerable from Library
 
hoelzl 
parents: 
63881 
diff
changeset
 | 
1  | 
(* Title: HOL/Analysis/Continuum_Not_Denumerable.thy  | 
| 56796 | 2  | 
Author: Benjamin Porter, Monash University, NICTA, 2005  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
3  | 
Author: Johannes Hölzl, TU München  | 
| 23461 | 4  | 
*)  | 
5  | 
||
| 69517 | 6  | 
section \<open>Non-Denumerability of the Continuum\<close>  | 
| 23461 | 7  | 
|
| 63464 | 8  | 
theory Continuum_Not_Denumerable  | 
| 
63970
 
3b6a3632e754
HOL-Analysis: move Continuum_Not_Denumerable from Library
 
hoelzl 
parents: 
63881 
diff
changeset
 | 
9  | 
imports  | 
| 
 
3b6a3632e754
HOL-Analysis: move Continuum_Not_Denumerable from Library
 
hoelzl 
parents: 
63881 
diff
changeset
 | 
10  | 
Complex_Main  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
63970 
diff
changeset
 | 
11  | 
"HOL-Library.Countable_Set"  | 
| 23461 | 12  | 
begin  | 
13  | 
||
| 70136 | 14  | 
subsection\<^marker>\<open>tag unimportant\<close> \<open>Abstract\<close>  | 
| 23461 | 15  | 
|
| 63465 | 16  | 
text \<open>  | 
17  | 
The following document presents a proof that the Continuum is uncountable.  | 
|
18  | 
It is formalised in the Isabelle/Isar theorem proving system.  | 
|
| 23461 | 19  | 
|
| 63465 | 20  | 
\<^bold>\<open>Theorem:\<close> The Continuum \<open>\<real>\<close> is not denumerable. In other words, there does  | 
| 63628 | 21  | 
not exist a function \<open>f: \<nat> \<Rightarrow> \<real>\<close> such that \<open>f\<close> is surjective.  | 
| 63465 | 22  | 
|
23  | 
\<^bold>\<open>Outline:\<close> An elegant informal proof of this result uses Cantor's  | 
|
24  | 
Diagonalisation argument. The proof presented here is not this one.  | 
|
| 23461 | 25  | 
|
| 63465 | 26  | 
First we formalise some properties of closed intervals, then we prove the  | 
27  | 
Nested Interval Property. This property relies on the completeness of the  | 
|
28  | 
Real numbers and is the foundation for our argument. Informally it states  | 
|
29  | 
that an intersection of countable closed intervals (where each successive  | 
|
30  | 
interval is a subset of the last) is non-empty. We then assume a surjective  | 
|
31  | 
function \<open>f: \<nat> \<Rightarrow> \<real>\<close> exists and find a real \<open>x\<close> such that \<open>x\<close> is not in the  | 
|
32  | 
range of \<open>f\<close> by generating a sequence of closed intervals then using the  | 
|
| 63628 | 33  | 
Nested Interval Property.  | 
| 63465 | 34  | 
\<close>  | 
| 70136 | 35  | 
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>  | 
| 
68607
 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 
immler 
parents: 
67968 
diff
changeset
 | 
36  | 
theorem real_non_denum: "\<nexists>f :: nat \<Rightarrow> real. surj f"  | 
| 
 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 
immler 
parents: 
67968 
diff
changeset
 | 
37  | 
proof  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
38  | 
assume "\<exists>f::nat \<Rightarrow> real. surj f"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
39  | 
then obtain f :: "nat \<Rightarrow> real" where "surj f" ..  | 
| 56796 | 40  | 
|
| 69597 | 41  | 
txt \<open>First we construct a sequence of nested intervals, ignoring \<^term>\<open>range f\<close>.\<close>  | 
| 23461 | 42  | 
|
| 63465 | 43  | 
  have "a < b \<Longrightarrow> \<exists>ka kb. ka < kb \<and> {ka..kb} \<subseteq> {a..b} \<and> c \<notin> {ka..kb}" for a b c :: real
 | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
44  | 
by (auto simp add: not_le cong: conj_cong)  | 
| 63465 | 45  | 
(metis dense le_less_linear less_linear less_trans order_refl)  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
46  | 
then obtain i j where ij:  | 
| 63465 | 47  | 
"a < b \<Longrightarrow> i a b c < j a b c"  | 
| 63060 | 48  | 
      "a < b \<Longrightarrow> {i a b c .. j a b c} \<subseteq> {a .. b}"
 | 
49  | 
      "a < b \<Longrightarrow> c \<notin> {i a b c .. j a b c}"
 | 
|
50  | 
for a b c :: real  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
51  | 
by metis  | 
| 23461 | 52  | 
|
| 63040 | 53  | 
define ivl where "ivl =  | 
54  | 
rec_nat (f 0 + 1, f 0 + 2) (\<lambda>n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))"  | 
|
55  | 
  define I where "I n = {fst (ivl n) .. snd (ivl n)}" for n
 | 
|
| 23461 | 56  | 
|
| 63465 | 57  | 
have ivl [simp]:  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
58  | 
"ivl 0 = (f 0 + 1, f 0 + 2)"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
59  | 
"\<And>n. ivl (Suc n) = (i (fst (ivl n)) (snd (ivl n)) (f n), j (fst (ivl n)) (snd (ivl n)) (f n))"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
60  | 
unfolding ivl_def by simp_all  | 
| 56796 | 61  | 
|
| 60500 | 62  | 
txt \<open>This is a decreasing sequence of non-empty intervals.\<close>  | 
| 23461 | 63  | 
|
| 63465 | 64  | 
have less: "fst (ivl n) < snd (ivl n)" for n  | 
65  | 
by (induct n) (auto intro!: ij)  | 
|
| 23461 | 66  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
67  | 
have "decseq I"  | 
| 63465 | 68  | 
unfolding I_def decseq_Suc_iff ivl fst_conv snd_conv  | 
69  | 
by (intro ij allI less)  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
70  | 
|
| 60500 | 71  | 
txt \<open>Now we apply the finite intersection property of compact sets.\<close>  | 
| 23461 | 72  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
73  | 
  have "I 0 \<inter> (\<Inter>i. I i) \<noteq> {}"
 | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
74  | 
proof (rule compact_imp_fip_image)  | 
| 63465 | 75  | 
fix S :: "nat set"  | 
76  | 
assume fin: "finite S"  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
77  | 
    have "{} \<subset> I (Max (insert 0 S))"
 | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
78  | 
unfolding I_def using less[of "Max (insert 0 S)"] by auto  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
79  | 
also have "I (Max (insert 0 S)) \<subseteq> (\<Inter>i\<in>insert 0 S. I i)"  | 
| 63465 | 80  | 
using fin decseqD[OF \<open>decseq I\<close>, of _ "Max (insert 0 S)"]  | 
81  | 
by (auto simp: Max_ge_iff)  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
82  | 
also have "(\<Inter>i\<in>insert 0 S. I i) = I 0 \<inter> (\<Inter>i\<in>S. I i)"  | 
| 56796 | 83  | 
by auto  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
84  | 
    finally show "I 0 \<inter> (\<Inter>i\<in>S. I i) \<noteq> {}"
 | 
| 53372 | 85  | 
by auto  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
86  | 
qed (auto simp: I_def)  | 
| 63060 | 87  | 
then obtain x where "x \<in> I n" for n  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
88  | 
by blast  | 
| 60500 | 89  | 
moreover from \<open>surj f\<close> obtain j where "x = f j"  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
90  | 
by blast  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
91  | 
ultimately have "f j \<in> I (Suc j)"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
92  | 
by blast  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
93  | 
with ij(3)[OF less] show False  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
94  | 
unfolding I_def ivl fst_conv snd_conv by auto  | 
| 23461 | 95  | 
qed  | 
96  | 
||
| 63465 | 97  | 
lemma uncountable_UNIV_real: "uncountable (UNIV :: real set)"  | 
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
98  | 
using real_non_denum unfolding uncountable_def by auto  | 
| 23461 | 99  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
100  | 
lemma bij_betw_open_intervals:  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
101  | 
fixes a b c d :: real  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
102  | 
assumes "a < b" "c < d"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
103  | 
  shows "\<exists>f. bij_betw f {a<..<b} {c<..<d}"
 | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
104  | 
proof -  | 
| 63040 | 105  | 
define f where "f a b c d x = (d - c)/(b - a) * (x - a) + c" for a b c d x :: real  | 
| 63465 | 106  | 
  {
 | 
107  | 
fix a b c d x :: real  | 
|
108  | 
assume *: "a < b" "c < d" "a < x" "x < b"  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
109  | 
moreover from * have "(d - c) * (x - a) < (d - c) * (b - a)"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
110  | 
by (intro mult_strict_left_mono) simp_all  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
111  | 
moreover from * have "0 < (d - c) * (x - a) / (b - a)"  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
112  | 
by simp  | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
113  | 
ultimately have "f a b c d x < d" "c < f a b c d x"  | 
| 63465 | 114  | 
by (simp_all add: f_def field_simps)  | 
115  | 
}  | 
|
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
116  | 
  with assms have "bij_betw (f a b c d) {a<..<b} {c<..<d}"
 | 
| 
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
117  | 
by (intro bij_betw_byWitness[where f'="f c d a b"]) (auto simp: f_def)  | 
| 63465 | 118  | 
then show ?thesis by auto  | 
| 23461 | 119  | 
qed  | 
120  | 
||
| 
57234
 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
56796 
diff
changeset
 | 
121  | 
lemma bij_betw_tan: "bij_betw tan {-pi/2<..<pi/2} UNIV"
 | 
| 
59872
 
db4000b71fdb
Theorem "arctan" is no longer a default simprule
 
paulson <lp15@cam.ac.uk> 
parents: 
59720 
diff
changeset
 | 
122  | 
using arctan_ubound by (intro bij_betw_byWitness[where f'=arctan]) (auto simp: arctan arctan_tan)  | 
| 23461 | 123  | 
|
| 63465 | 124  | 
lemma uncountable_open_interval: "uncountable {a<..<b} \<longleftrightarrow> a < b" for a b :: real
 | 
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
125  | 
proof  | 
| 63465 | 126  | 
  show "a < b" if "uncountable {a<..<b}"
 | 
127  | 
using uncountable_def that by force  | 
|
128  | 
  show "uncountable {a<..<b}" if "a < b"
 | 
|
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
129  | 
proof -  | 
| 
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
130  | 
    obtain f where "bij_betw f {a <..< b} {-pi/2<..<pi/2}"
 | 
| 60500 | 131  | 
using bij_betw_open_intervals[OF \<open>a < b\<close>, of "-pi/2" "pi/2"] by auto  | 
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
132  | 
then show ?thesis  | 
| 
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
133  | 
by (metis bij_betw_tan uncountable_bij_betw uncountable_UNIV_real)  | 
| 
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
134  | 
qed  | 
| 23461 | 135  | 
qed  | 
136  | 
||
| 63465 | 137  | 
lemma uncountable_half_open_interval_1: "uncountable {a..<b} \<longleftrightarrow> a < b" for a b :: real
 | 
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
138  | 
apply auto  | 
| 63465 | 139  | 
using atLeastLessThan_empty_iff  | 
140  | 
apply fastforce  | 
|
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
141  | 
using uncountable_open_interval [of a b]  | 
| 63465 | 142  | 
apply (metis countable_Un_iff ivl_disj_un_singleton(3))  | 
143  | 
done  | 
|
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
144  | 
|
| 63465 | 145  | 
lemma uncountable_half_open_interval_2: "uncountable {a<..b} \<longleftrightarrow> a < b" for a b :: real
 | 
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
146  | 
apply auto  | 
| 63465 | 147  | 
using atLeastLessThan_empty_iff  | 
148  | 
apply fastforce  | 
|
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
149  | 
using uncountable_open_interval [of a b]  | 
| 63465 | 150  | 
apply (metis countable_Un_iff ivl_disj_un_singleton(4))  | 
151  | 
done  | 
|
| 
60308
 
f7e406aba90d
uncountability: open interval equivalences
 
paulson <lp15@cam.ac.uk> 
parents: 
59872 
diff
changeset
 | 
152  | 
|
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
153  | 
lemma real_interval_avoid_countable_set:  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
154  | 
fixes a b :: real and A :: "real set"  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
155  | 
assumes "a < b" and "countable A"  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
156  | 
  shows "\<exists>x\<in>{a<..<b}. x \<notin> A"
 | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
157  | 
proof -  | 
| 63540 | 158  | 
  from \<open>countable A\<close> have *: "countable (A \<inter> {a<..<b})"
 | 
| 63465 | 159  | 
by auto  | 
| 63540 | 160  | 
  with \<open>a < b\<close> have "\<not> countable {a<..<b}"
 | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
161  | 
by (simp add: uncountable_open_interval)  | 
| 63540 | 162  | 
  with * have "A \<inter> {a<..<b} \<noteq> {a<..<b}"
 | 
| 63465 | 163  | 
by auto  | 
164  | 
  then have "A \<inter> {a<..<b} \<subset> {a<..<b}"
 | 
|
165  | 
by (intro psubsetI) auto  | 
|
166  | 
  then have "\<exists>x. x \<in> {a<..<b} - A \<inter> {a<..<b}"
 | 
|
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
167  | 
by (rule psubset_imp_ex_mem)  | 
| 63465 | 168  | 
then show ?thesis  | 
169  | 
by auto  | 
|
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
170  | 
qed  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61585 
diff
changeset
 | 
171  | 
|
| 
63881
 
b746b19197bd
lots of new results about topology, affine dimension etc
 
paulson <lp15@cam.ac.uk> 
parents: 
63628 
diff
changeset
 | 
172  | 
lemma uncountable_closed_interval: "uncountable {a..b} \<longleftrightarrow> a < b" for a b :: real
 | 
| 
 
b746b19197bd
lots of new results about topology, affine dimension etc
 
paulson <lp15@cam.ac.uk> 
parents: 
63628 
diff
changeset
 | 
173  | 
apply (rule iffI)  | 
| 
 
b746b19197bd
lots of new results about topology, affine dimension etc
 
paulson <lp15@cam.ac.uk> 
parents: 
63628 
diff
changeset
 | 
174  | 
apply (metis atLeastAtMost_singleton atLeastatMost_empty countable_finite finite.emptyI finite_insert linorder_neqE_linordered_idom)  | 
| 
 
b746b19197bd
lots of new results about topology, affine dimension etc
 
paulson <lp15@cam.ac.uk> 
parents: 
63628 
diff
changeset
 | 
175  | 
using real_interval_avoid_countable_set by fastforce  | 
| 
 
b746b19197bd
lots of new results about topology, affine dimension etc
 
paulson <lp15@cam.ac.uk> 
parents: 
63628 
diff
changeset
 | 
176  | 
|
| 62083 | 177  | 
lemma open_minus_countable:  | 
| 63465 | 178  | 
fixes S A :: "real set"  | 
179  | 
  assumes "countable A" "S \<noteq> {}" "open S"
 | 
|
| 62083 | 180  | 
shows "\<exists>x\<in>S. x \<notin> A"  | 
181  | 
proof -  | 
|
182  | 
obtain x where "x \<in> S"  | 
|
183  | 
    using \<open>S \<noteq> {}\<close> by auto
 | 
|
184  | 
  then obtain e where "0 < e" "{y. dist y x < e} \<subseteq> S"
 | 
|
185  | 
using \<open>open S\<close> by (auto simp: open_dist subset_eq)  | 
|
186  | 
  moreover have "{y. dist y x < e} = {x - e <..< x + e}"
 | 
|
187  | 
by (auto simp: dist_real_def)  | 
|
188  | 
ultimately have "uncountable (S - A)"  | 
|
189  | 
using uncountable_open_interval[of "x - e" "x + e"] \<open>countable A\<close>  | 
|
190  | 
by (intro uncountable_minus_countable) (auto dest: countable_subset)  | 
|
191  | 
then show ?thesis  | 
|
192  | 
unfolding uncountable_def by auto  | 
|
193  | 
qed  | 
|
194  | 
||
| 23461 | 195  | 
end  |