author | haftmann |
Thu, 26 Jun 2008 10:07:01 +0200 | |
changeset 27368 | 9f90ac19e32b |
parent 27110 | 194aa674c2a1 |
child 27487 | c8a6ce181805 |
permissions | -rw-r--r-- |
11355 | 1 |
(* Title: HOL/Library/Nat_Infinity.thy |
2 |
ID: $Id$ |
|
27110 | 3 |
Author: David von Oheimb, TU Muenchen; Florian Haftmann, TU Muenchen |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
4 |
*) |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
5 |
|
14706 | 6 |
header {* Natural numbers with infinity *} |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
7 |
|
15131 | 8 |
theory Nat_Infinity |
27368 | 9 |
imports Plain Presburger |
15131 | 10 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
11 |
|
27110 | 12 |
subsection {* Type definition *} |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
13 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
14 |
text {* |
11355 | 15 |
We extend the standard natural numbers by a special value indicating |
27110 | 16 |
infinity. |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
17 |
*} |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
18 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
19 |
datatype inat = Fin nat | Infty |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
20 |
|
21210 | 21 |
notation (xsymbols) |
19736 | 22 |
Infty ("\<infinity>") |
23 |
||
21210 | 24 |
notation (HTML output) |
19736 | 25 |
Infty ("\<infinity>") |
26 |
||
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
27 |
|
27110 | 28 |
subsection {* Constructors and numbers *} |
29 |
||
30 |
instantiation inat :: "{zero, one, number}" |
|
25594 | 31 |
begin |
32 |
||
33 |
definition |
|
27110 | 34 |
"0 = Fin 0" |
25594 | 35 |
|
36 |
definition |
|
27110 | 37 |
[code inline]: "1 = Fin 1" |
25594 | 38 |
|
39 |
definition |
|
27110 | 40 |
[code inline, code func del]: "number_of k = Fin (number_of k)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
41 |
|
25594 | 42 |
instance .. |
43 |
||
44 |
end |
|
45 |
||
27110 | 46 |
definition iSuc :: "inat \<Rightarrow> inat" where |
47 |
"iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)" |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
48 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
49 |
lemma Fin_0: "Fin 0 = 0" |
27110 | 50 |
by (simp add: zero_inat_def) |
51 |
||
52 |
lemma Fin_1: "Fin 1 = 1" |
|
53 |
by (simp add: one_inat_def) |
|
54 |
||
55 |
lemma Fin_number: "Fin (number_of k) = number_of k" |
|
56 |
by (simp add: number_of_inat_def) |
|
57 |
||
58 |
lemma one_iSuc: "1 = iSuc 0" |
|
59 |
by (simp add: zero_inat_def one_inat_def iSuc_def) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
60 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
61 |
lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0" |
27110 | 62 |
by (simp add: zero_inat_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
63 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
64 |
lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>" |
27110 | 65 |
by (simp add: zero_inat_def) |
66 |
||
67 |
lemma zero_inat_eq [simp]: |
|
68 |
"number_of k = (0\<Colon>inat) \<longleftrightarrow> number_of k = (0\<Colon>nat)" |
|
69 |
"(0\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)" |
|
70 |
unfolding zero_inat_def number_of_inat_def by simp_all |
|
71 |
||
72 |
lemma one_inat_eq [simp]: |
|
73 |
"number_of k = (1\<Colon>inat) \<longleftrightarrow> number_of k = (1\<Colon>nat)" |
|
74 |
"(1\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)" |
|
75 |
unfolding one_inat_def number_of_inat_def by simp_all |
|
76 |
||
77 |
lemma zero_one_inat_neq [simp]: |
|
78 |
"\<not> 0 = (1\<Colon>inat)" |
|
79 |
"\<not> 1 = (0\<Colon>inat)" |
|
80 |
unfolding zero_inat_def one_inat_def by simp_all |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
81 |
|
27110 | 82 |
lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1" |
83 |
by (simp add: one_inat_def) |
|
84 |
||
85 |
lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>" |
|
86 |
by (simp add: one_inat_def) |
|
87 |
||
88 |
lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k" |
|
89 |
by (simp add: number_of_inat_def) |
|
90 |
||
91 |
lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>" |
|
92 |
by (simp add: number_of_inat_def) |
|
93 |
||
94 |
lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)" |
|
95 |
by (simp add: iSuc_def) |
|
96 |
||
97 |
lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))" |
|
98 |
by (simp add: iSuc_Fin number_of_inat_def) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
99 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
100 |
lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>" |
27110 | 101 |
by (simp add: iSuc_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
102 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
103 |
lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0" |
27110 | 104 |
by (simp add: iSuc_def zero_inat_def split: inat.splits) |
105 |
||
106 |
lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n" |
|
107 |
by (rule iSuc_ne_0 [symmetric]) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
108 |
|
27110 | 109 |
lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n" |
110 |
by (simp add: iSuc_def split: inat.splits) |
|
111 |
||
112 |
lemma number_of_inat_inject [simp]: |
|
113 |
"(number_of k \<Colon> inat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l" |
|
114 |
by (simp add: number_of_inat_def) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
115 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
116 |
|
27110 | 117 |
subsection {* Addition *} |
118 |
||
119 |
instantiation inat :: comm_monoid_add |
|
120 |
begin |
|
121 |
||
122 |
definition |
|
123 |
[code del]: "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))" |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
124 |
|
27110 | 125 |
lemma plus_inat_simps [simp, code]: |
126 |
"Fin m + Fin n = Fin (m + n)" |
|
127 |
"\<infinity> + q = \<infinity>" |
|
128 |
"q + \<infinity> = \<infinity>" |
|
129 |
by (simp_all add: plus_inat_def split: inat.splits) |
|
130 |
||
131 |
instance proof |
|
132 |
fix n m q :: inat |
|
133 |
show "n + m + q = n + (m + q)" |
|
134 |
by (cases n, auto, cases m, auto, cases q, auto) |
|
135 |
show "n + m = m + n" |
|
136 |
by (cases n, auto, cases m, auto) |
|
137 |
show "0 + n = n" |
|
138 |
by (cases n) (simp_all add: zero_inat_def) |
|
26089 | 139 |
qed |
140 |
||
27110 | 141 |
end |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
142 |
|
27110 | 143 |
lemma plus_inat_0 [simp]: |
144 |
"0 + (q\<Colon>inat) = q" |
|
145 |
"(q\<Colon>inat) + 0 = q" |
|
146 |
by (simp_all add: plus_inat_def zero_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
147 |
|
27110 | 148 |
lemma plus_inat_number [simp]: |
149 |
"(number_of k \<Colon> inat) + number_of l = (if neg (number_of k \<Colon> int) then number_of l |
|
150 |
else if neg (number_of l \<Colon> int) then number_of k else number_of (k + l))" |
|
151 |
unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] .. |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
152 |
|
27110 | 153 |
lemma iSuc_number [simp]: |
154 |
"iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))" |
|
155 |
unfolding iSuc_number_of |
|
156 |
unfolding one_inat_def number_of_inat_def Suc_nat_number_of if_distrib [symmetric] .. |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
157 |
|
27110 | 158 |
lemma iSuc_plus_1: |
159 |
"iSuc n = n + 1" |
|
160 |
by (cases n) (simp_all add: iSuc_Fin one_inat_def) |
|
161 |
||
162 |
lemma plus_1_iSuc: |
|
163 |
"1 + q = iSuc q" |
|
164 |
"q + 1 = iSuc q" |
|
165 |
unfolding iSuc_plus_1 by (simp_all add: add_ac) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
166 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
167 |
|
27110 | 168 |
subsection {* Ordering *} |
169 |
||
170 |
instantiation inat :: ordered_ab_semigroup_add |
|
171 |
begin |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
172 |
|
27110 | 173 |
definition |
174 |
[code del]: "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False) |
|
175 |
| \<infinity> \<Rightarrow> True)" |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
176 |
|
27110 | 177 |
definition |
178 |
[code del]: "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True) |
|
179 |
| \<infinity> \<Rightarrow> False)" |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
180 |
|
27110 | 181 |
lemma inat_ord_simps [simp]: |
182 |
"Fin m \<le> Fin n \<longleftrightarrow> m \<le> n" |
|
183 |
"Fin m < Fin n \<longleftrightarrow> m < n" |
|
184 |
"q \<le> \<infinity>" |
|
185 |
"q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>" |
|
186 |
"\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>" |
|
187 |
"\<infinity> < q \<longleftrightarrow> False" |
|
188 |
by (simp_all add: less_eq_inat_def less_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
189 |
|
27110 | 190 |
lemma inat_ord_code [code]: |
191 |
"Fin m \<le> Fin n \<longleftrightarrow> m \<le> n" |
|
192 |
"Fin m < Fin n \<longleftrightarrow> m < n" |
|
193 |
"q \<le> \<infinity> \<longleftrightarrow> True" |
|
194 |
"Fin m < \<infinity> \<longleftrightarrow> True" |
|
195 |
"\<infinity> \<le> Fin n \<longleftrightarrow> False" |
|
196 |
"\<infinity> < q \<longleftrightarrow> False" |
|
197 |
by simp_all |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
198 |
|
27110 | 199 |
instance by default |
200 |
(auto simp add: less_eq_inat_def less_inat_def plus_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
201 |
|
27110 | 202 |
end |
203 |
||
204 |
lemma inat_ord_number [simp]: |
|
205 |
"(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n" |
|
206 |
"(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n" |
|
207 |
by (simp_all add: number_of_inat_def) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
208 |
|
27110 | 209 |
lemma i0_lb [simp]: "(0\<Colon>inat) \<le> n" |
210 |
by (simp add: zero_inat_def less_eq_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
211 |
|
27110 | 212 |
lemma i0_neq [simp]: "n \<le> (0\<Colon>inat) \<longleftrightarrow> n = 0" |
213 |
by (simp add: zero_inat_def less_eq_inat_def split: inat.splits) |
|
214 |
||
215 |
lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R" |
|
216 |
by (simp add: zero_inat_def less_eq_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
217 |
|
27110 | 218 |
lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R" |
219 |
by simp |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
220 |
|
27110 | 221 |
lemma not_ilessi0 [simp]: "\<not> n < (0\<Colon>inat)" |
222 |
by (simp add: zero_inat_def less_inat_def split: inat.splits) |
|
223 |
||
224 |
lemma i0_eq [simp]: "(0\<Colon>inat) < n \<longleftrightarrow> n \<noteq> 0" |
|
225 |
by (simp add: zero_inat_def less_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
226 |
|
27110 | 227 |
lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m" |
228 |
by (simp add: iSuc_def less_eq_inat_def split: inat.splits) |
|
229 |
||
230 |
lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m" |
|
231 |
by (simp add: iSuc_def less_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
232 |
|
27110 | 233 |
lemma ile_iSuc [simp]: "n \<le> iSuc n" |
234 |
by (simp add: iSuc_def less_eq_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
235 |
|
11355 | 236 |
lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0" |
27110 | 237 |
by (simp add: zero_inat_def iSuc_def less_eq_inat_def split: inat.splits) |
238 |
||
239 |
lemma i0_iless_iSuc [simp]: "0 < iSuc n" |
|
240 |
by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.splits) |
|
241 |
||
242 |
lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n" |
|
243 |
by (simp add: iSuc_def less_eq_inat_def less_inat_def split: inat.splits) |
|
244 |
||
245 |
lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n" |
|
246 |
by (cases n) auto |
|
247 |
||
248 |
lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n" |
|
249 |
by (auto simp add: iSuc_def less_inat_def split: inat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
250 |
|
27110 | 251 |
lemma min_inat_simps [simp]: |
252 |
"min (Fin m) (Fin n) = Fin (min m n)" |
|
253 |
"min q 0 = 0" |
|
254 |
"min 0 q = 0" |
|
255 |
"min q \<infinity> = q" |
|
256 |
"min \<infinity> q = q" |
|
257 |
by (auto simp add: min_def) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
258 |
|
27110 | 259 |
lemma max_inat_simps [simp]: |
260 |
"max (Fin m) (Fin n) = Fin (max m n)" |
|
261 |
"max q 0 = q" |
|
262 |
"max 0 q = q" |
|
263 |
"max q \<infinity> = \<infinity>" |
|
264 |
"max \<infinity> q = \<infinity>" |
|
265 |
by (simp_all add: max_def) |
|
266 |
||
267 |
lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k" |
|
268 |
by (cases n) simp_all |
|
269 |
||
270 |
lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k" |
|
271 |
by (cases n) simp_all |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
272 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
273 |
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
274 |
apply (induct_tac k) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
275 |
apply (simp (no_asm) only: Fin_0) |
27110 | 276 |
apply (fast intro: le_less_trans [OF i0_lb]) |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
277 |
apply (erule exE) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
278 |
apply (drule spec) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
279 |
apply (erule exE) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
280 |
apply (drule ileI1) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
281 |
apply (rule iSuc_Fin [THEN subst]) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
282 |
apply (rule exI) |
27110 | 283 |
apply (erule (1) le_less_trans) |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
284 |
done |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
285 |
|
26089 | 286 |
|
27110 | 287 |
subsection {* Well-ordering *} |
26089 | 288 |
|
289 |
lemma less_FinE: |
|
290 |
"[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P" |
|
291 |
by (induct n) auto |
|
292 |
||
293 |
lemma less_InftyE: |
|
294 |
"[| n < Infty; !!k. n = Fin k ==> P |] ==> P" |
|
295 |
by (induct n) auto |
|
296 |
||
297 |
lemma inat_less_induct: |
|
298 |
assumes prem: "!!n. \<forall>m::inat. m < n --> P m ==> P n" shows "P n" |
|
299 |
proof - |
|
300 |
have P_Fin: "!!k. P (Fin k)" |
|
301 |
apply (rule nat_less_induct) |
|
302 |
apply (rule prem, clarify) |
|
303 |
apply (erule less_FinE, simp) |
|
304 |
done |
|
305 |
show ?thesis |
|
306 |
proof (induct n) |
|
307 |
fix nat |
|
308 |
show "P (Fin nat)" by (rule P_Fin) |
|
309 |
next |
|
310 |
show "P Infty" |
|
311 |
apply (rule prem, clarify) |
|
312 |
apply (erule less_InftyE) |
|
313 |
apply (simp add: P_Fin) |
|
314 |
done |
|
315 |
qed |
|
316 |
qed |
|
317 |
||
318 |
instance inat :: wellorder |
|
319 |
proof |
|
320 |
show "wf {(x::inat, y::inat). x < y}" |
|
321 |
proof (rule wfUNIVI) |
|
322 |
fix P and x :: inat |
|
323 |
assume "\<forall>x::inat. (\<forall>y. (y, x) \<in> {(x, y). x < y} \<longrightarrow> P y) \<longrightarrow> P x" |
|
324 |
hence 1: "!!x::inat. ALL y. y < x --> P y ==> P x" by fast |
|
325 |
thus "P x" by (rule inat_less_induct) |
|
326 |
qed |
|
327 |
qed |
|
328 |
||
27110 | 329 |
|
330 |
subsection {* Traditional theorem names *} |
|
331 |
||
332 |
lemmas inat_defs = zero_inat_def one_inat_def number_of_inat_def iSuc_def |
|
333 |
plus_inat_def less_eq_inat_def less_inat_def |
|
334 |
||
335 |
lemmas inat_splits = inat.splits |
|
336 |
||
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
337 |
end |