author | wenzelm |
Wed, 04 May 2011 15:37:39 +0200 | |
changeset 42676 | 8724f20bf69c |
parent 37388 | 793618618f78 |
permissions | -rw-r--r-- |
30326
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
1 |
(* Author: Florian Haftmann, TU Muenchen *) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
2 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
3 |
header {* A constructive version of Cantor's first diagonalization argument. *} |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
4 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
5 |
theory Diagonalize |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
6 |
imports Main |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
7 |
begin |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
8 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
9 |
subsection {* Summation from @{text "0"} to @{text "n"} *} |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
10 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
11 |
definition sum :: "nat \<Rightarrow> nat" where |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
12 |
"sum n = n * Suc n div 2" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
13 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
14 |
lemma sum_0: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
15 |
"sum 0 = 0" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
16 |
unfolding sum_def by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
17 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
18 |
lemma sum_Suc: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
19 |
"sum (Suc n) = Suc n + sum n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
20 |
unfolding sum_def by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
21 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
22 |
lemma sum2: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
23 |
"2 * sum n = n * Suc n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
24 |
proof - |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
25 |
have "2 dvd n * Suc n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
26 |
proof (cases "2 dvd n") |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
27 |
case True then show ?thesis by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
28 |
next |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
29 |
case False then have "2 dvd Suc n" by arith |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
30 |
then show ?thesis by (simp del: mult_Suc_right) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
31 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
32 |
then have "n * Suc n div 2 * 2 = n * Suc n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
33 |
by (rule dvd_div_mult_self [of "2::nat"]) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
34 |
then show ?thesis by (simp add: sum_def) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
35 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
36 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
37 |
lemma sum_strict_mono: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
38 |
"strict_mono sum" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
39 |
proof (rule strict_monoI) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
40 |
fix m n :: nat |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
41 |
assume "m < n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
42 |
then have "m * Suc m < n * Suc n" by (intro mult_strict_mono) simp_all |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
43 |
then have "2 * sum m < 2 * sum n" by (simp add: sum2) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
44 |
then show "sum m < sum n" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
45 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
46 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
47 |
lemma sum_not_less_self: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
48 |
"n \<le> sum n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
49 |
proof - |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
50 |
have "2 * n \<le> n * Suc n" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
51 |
with sum2 have "2 * n \<le> 2 * sum n" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
52 |
then show ?thesis by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
53 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
54 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
55 |
lemma sum_rest_aux: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
56 |
assumes "q \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
57 |
assumes "sum m \<le> sum n + q" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
58 |
shows "m \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
59 |
proof (rule ccontr) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
60 |
assume "\<not> m \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
61 |
then have "n < m" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
62 |
then have "m \<ge> Suc n" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
63 |
then have "m = m - Suc n + Suc n" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
64 |
then have "m = Suc (n + (m - Suc n))" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
65 |
then obtain r where "m = Suc (n + r)" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
66 |
with `sum m \<le> sum n + q` have "sum (Suc (n + r)) \<le> sum n + q" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
67 |
then have "sum (n + r) + Suc (n + r) \<le> sum n + q" unfolding sum_Suc by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
68 |
with `m = Suc (n + r)` have "sum (n + r) + m \<le> sum n + q" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
69 |
have "sum n \<le> sum (n + r)" unfolding strict_mono_less_eq [OF sum_strict_mono] by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
70 |
moreover from `q \<le> n` `n < m` have "q < m" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
71 |
ultimately have "sum n + q < sum (n + r) + m" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
72 |
with `sum (n + r) + m \<le> sum n + q` show False |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
73 |
by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
74 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
75 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
76 |
lemma sum_rest: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
77 |
assumes "q \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
78 |
shows "sum m \<le> sum n + q \<longleftrightarrow> m \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
79 |
using assms apply (auto intro: sum_rest_aux) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
80 |
apply (simp add: strict_mono_less_eq [OF sum_strict_mono, symmetric, of m n]) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
81 |
done |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
82 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
83 |
|
37388 | 84 |
subsection {* Diagonalization: an injective embedding of two @{typ nat}s to one @{typ nat} *} |
30326
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
85 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
86 |
definition diagonalize :: "nat \<Rightarrow> nat \<Rightarrow> nat" where |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
87 |
"diagonalize m n = sum (m + n) + m" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
88 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
89 |
lemma diagonalize_inject: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
90 |
assumes "diagonalize a b = diagonalize c d" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
91 |
shows "a = c" and "b = d" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
92 |
proof - |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
93 |
from assms have diageq: "sum (a + b) + a = sum (c + d) + c" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
94 |
by (simp add: diagonalize_def) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
95 |
have "a + b = c + d \<or> a + b \<ge> Suc (c + d) \<or> c + d \<ge> Suc (a + b)" by arith |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
96 |
then have "a = c \<and> b = d" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
97 |
proof (elim disjE) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
98 |
assume sumeq: "a + b = c + d" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
99 |
then have "a = c" using diageq by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
100 |
moreover from sumeq this have "b = d" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
101 |
ultimately show ?thesis .. |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
102 |
next |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
103 |
assume "a + b \<ge> Suc (c + d)" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
104 |
with strict_mono_less_eq [OF sum_strict_mono] |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
105 |
have "sum (a + b) \<ge> sum (Suc (c + d))" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
106 |
with diageq show ?thesis by (simp add: sum_Suc) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
107 |
next |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
108 |
assume "c + d \<ge> Suc (a + b)" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
109 |
with strict_mono_less_eq [OF sum_strict_mono] |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
110 |
have "sum (c + d) \<ge> sum (Suc (a + b))" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
111 |
with diageq show ?thesis by (simp add: sum_Suc) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
112 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
113 |
then show "a = c" and "b = d" by auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
114 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
115 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
116 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
117 |
subsection {* The reverse diagonalization: reconstruction a pair of @{typ nat}s from one @{typ nat} *} |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
118 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
119 |
text {* The inverse of the @{const sum} function *} |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
120 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
121 |
definition tupelize :: "nat \<Rightarrow> nat \<times> nat" where |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
122 |
"tupelize q = (let d = Max {d. sum d \<le> q}; m = q - sum d |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
123 |
in (m, d - m))" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
124 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
125 |
lemma tupelize_diagonalize: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
126 |
"tupelize (diagonalize m n) = (m, n)" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
127 |
proof - |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
128 |
from sum_rest |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
129 |
have "\<And>r. sum r \<le> sum (m + n) + m \<longleftrightarrow> r \<le> m + n" by simp |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
130 |
then have "Max {d. sum d \<le> (sum (m + n) + m)} = m + n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
131 |
by (auto intro: Max_eqI) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
132 |
then show ?thesis |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
133 |
by (simp add: tupelize_def diagonalize_def) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
134 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
135 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
136 |
lemma snd_tupelize: |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
137 |
"snd (tupelize n) \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
138 |
proof - |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
139 |
have "sum 0 \<le> n" by (simp add: sum_0) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
140 |
then have "Max {m \<Colon> nat. sum m \<le> n} \<le> Max {m \<Colon> nat. m \<le> n}" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
141 |
by (intro Max_mono [of "{m. sum m \<le> n}" "{m. m \<le> n}"]) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
142 |
(auto intro: Max_mono order_trans sum_not_less_self) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
143 |
also have "Max {m \<Colon> nat. m \<le> n} \<le> n" |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
144 |
by (subst Max_le_iff) auto |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
145 |
finally have "Max {m. sum m \<le> n} \<le> n" . |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
146 |
then show ?thesis by (simp add: tupelize_def Let_def) |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
147 |
qed |
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
148 |
|
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
haftmann
parents:
diff
changeset
|
149 |
end |