author | wenzelm |
Fri, 03 Jul 1998 17:34:24 +0200 | |
changeset 5122 | 229190f9f303 |
parent 5069 | 3ea049f7979d |
child 5132 | 24f992a25adc |
permissions | -rw-r--r-- |
2608 | 1 |
(* Title: HOL/NatDef.ML |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
|
4 |
Copyright 1991 University of Cambridge |
|
4737 | 5 |
|
6 |
Blast_tac proofs here can get PROOF FAILED of Ord theorems like order_refl |
|
7 |
and order_less_irrefl. We do not add the "nat" versions to the basic claset |
|
8 |
because the type will be promoted to type class "order". |
|
2608 | 9 |
*) |
10 |
||
5069 | 11 |
Goal "mono(%X. {Zero_Rep} Un (Suc_Rep``X))"; |
2608 | 12 |
by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1)); |
13 |
qed "Nat_fun_mono"; |
|
14 |
||
15 |
val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski); |
|
16 |
||
17 |
(* Zero is a natural number -- this also justifies the type definition*) |
|
5069 | 18 |
Goal "Zero_Rep: Nat"; |
2608 | 19 |
by (stac Nat_unfold 1); |
20 |
by (rtac (singletonI RS UnI1) 1); |
|
21 |
qed "Zero_RepI"; |
|
22 |
||
23 |
val prems = goal thy "i: Nat ==> Suc_Rep(i) : Nat"; |
|
24 |
by (stac Nat_unfold 1); |
|
25 |
by (rtac (imageI RS UnI2) 1); |
|
26 |
by (resolve_tac prems 1); |
|
27 |
qed "Suc_RepI"; |
|
28 |
||
29 |
(*** Induction ***) |
|
30 |
||
31 |
val major::prems = goal thy |
|
32 |
"[| i: Nat; P(Zero_Rep); \ |
|
33 |
\ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)"; |
|
34 |
by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); |
|
4089 | 35 |
by (blast_tac (claset() addIs prems) 1); |
2608 | 36 |
qed "Nat_induct"; |
37 |
||
38 |
val prems = goalw thy [Zero_def,Suc_def] |
|
39 |
"[| P(0); \ |
|
3040
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
3023
diff
changeset
|
40 |
\ !!n. P(n) ==> P(Suc(n)) |] ==> P(n)"; |
2608 | 41 |
by (rtac (Rep_Nat_inverse RS subst) 1); (*types force good instantiation*) |
42 |
by (rtac (Rep_Nat RS Nat_induct) 1); |
|
43 |
by (REPEAT (ares_tac prems 1 |
|
44 |
ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1)); |
|
45 |
qed "nat_induct"; |
|
46 |
||
47 |
(*Perform induction on n. *) |
|
3563 | 48 |
local fun raw_nat_ind_tac a i = |
49 |
res_inst_tac [("n",a)] nat_induct i THEN rename_last_tac a [""] (i+1) |
|
50 |
in |
|
51 |
val nat_ind_tac = Datatype.occs_in_prems raw_nat_ind_tac |
|
52 |
end; |
|
3040
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
3023
diff
changeset
|
53 |
|
2608 | 54 |
(*A special form of induction for reasoning about m<n and m-n*) |
55 |
val prems = goal thy |
|
56 |
"[| !!x. P x 0; \ |
|
57 |
\ !!y. P 0 (Suc y); \ |
|
58 |
\ !!x y. [| P x y |] ==> P (Suc x) (Suc y) \ |
|
59 |
\ |] ==> P m n"; |
|
60 |
by (res_inst_tac [("x","m")] spec 1); |
|
61 |
by (nat_ind_tac "n" 1); |
|
62 |
by (rtac allI 2); |
|
63 |
by (nat_ind_tac "x" 2); |
|
64 |
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1)); |
|
65 |
qed "diff_induct"; |
|
66 |
||
67 |
(*Case analysis on the natural numbers*) |
|
68 |
val prems = goal thy |
|
69 |
"[| n=0 ==> P; !!x. n = Suc(x) ==> P |] ==> P"; |
|
70 |
by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1); |
|
4089 | 71 |
by (fast_tac (claset() addSEs prems) 1); |
2608 | 72 |
by (nat_ind_tac "n" 1); |
73 |
by (rtac (refl RS disjI1) 1); |
|
2891 | 74 |
by (Blast_tac 1); |
2608 | 75 |
qed "natE"; |
76 |
||
3282
c31e6239d4c9
Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents:
3236
diff
changeset
|
77 |
|
2608 | 78 |
(*** Isomorphisms: Abs_Nat and Rep_Nat ***) |
79 |
||
80 |
(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat), |
|
81 |
since we assume the isomorphism equations will one day be given by Isabelle*) |
|
82 |
||
5069 | 83 |
Goal "inj(Rep_Nat)"; |
2608 | 84 |
by (rtac inj_inverseI 1); |
85 |
by (rtac Rep_Nat_inverse 1); |
|
86 |
qed "inj_Rep_Nat"; |
|
87 |
||
5069 | 88 |
Goal "inj_on Abs_Nat Nat"; |
4830 | 89 |
by (rtac inj_on_inverseI 1); |
2608 | 90 |
by (etac Abs_Nat_inverse 1); |
4830 | 91 |
qed "inj_on_Abs_Nat"; |
2608 | 92 |
|
93 |
(*** Distinctness of constructors ***) |
|
94 |
||
5069 | 95 |
Goalw [Zero_def,Suc_def] "Suc(m) ~= 0"; |
4830 | 96 |
by (rtac (inj_on_Abs_Nat RS inj_on_contraD) 1); |
2608 | 97 |
by (rtac Suc_Rep_not_Zero_Rep 1); |
98 |
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); |
|
99 |
qed "Suc_not_Zero"; |
|
100 |
||
101 |
bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym); |
|
102 |
||
103 |
AddIffs [Suc_not_Zero,Zero_not_Suc]; |
|
104 |
||
105 |
bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE)); |
|
106 |
val Zero_neq_Suc = sym RS Suc_neq_Zero; |
|
107 |
||
108 |
(** Injectiveness of Suc **) |
|
109 |
||
5069 | 110 |
Goalw [Suc_def] "inj(Suc)"; |
2608 | 111 |
by (rtac injI 1); |
4830 | 112 |
by (dtac (inj_on_Abs_Nat RS inj_onD) 1); |
2608 | 113 |
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); |
114 |
by (dtac (inj_Suc_Rep RS injD) 1); |
|
115 |
by (etac (inj_Rep_Nat RS injD) 1); |
|
116 |
qed "inj_Suc"; |
|
117 |
||
118 |
val Suc_inject = inj_Suc RS injD; |
|
119 |
||
5069 | 120 |
Goal "(Suc(m)=Suc(n)) = (m=n)"; |
2608 | 121 |
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); |
122 |
qed "Suc_Suc_eq"; |
|
123 |
||
124 |
AddIffs [Suc_Suc_eq]; |
|
125 |
||
5069 | 126 |
Goal "n ~= Suc(n)"; |
2608 | 127 |
by (nat_ind_tac "n" 1); |
128 |
by (ALLGOALS Asm_simp_tac); |
|
129 |
qed "n_not_Suc_n"; |
|
130 |
||
131 |
bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym); |
|
132 |
||
5069 | 133 |
Goal "!!n. n ~= 0 ==> EX m. n = Suc m"; |
3457 | 134 |
by (rtac natE 1); |
3236 | 135 |
by (REPEAT (Blast_tac 1)); |
136 |
qed "not0_implies_Suc"; |
|
137 |
||
138 |
||
2608 | 139 |
(*** nat_case -- the selection operator for nat ***) |
140 |
||
5069 | 141 |
Goalw [nat_case_def] "nat_case a f 0 = a"; |
4535 | 142 |
by (Blast_tac 1); |
2608 | 143 |
qed "nat_case_0"; |
144 |
||
5069 | 145 |
Goalw [nat_case_def] "nat_case a f (Suc k) = f(k)"; |
4535 | 146 |
by (Blast_tac 1); |
2608 | 147 |
qed "nat_case_Suc"; |
148 |
||
5069 | 149 |
Goalw [wf_def, pred_nat_def] "wf(pred_nat)"; |
3718 | 150 |
by (Clarify_tac 1); |
2608 | 151 |
by (nat_ind_tac "x" 1); |
3236 | 152 |
by (ALLGOALS Blast_tac); |
2608 | 153 |
qed "wf_pred_nat"; |
154 |
||
155 |
||
156 |
(*** nat_rec -- by wf recursion on pred_nat ***) |
|
157 |
||
158 |
(* The unrolling rule for nat_rec *) |
|
5069 | 159 |
Goal |
4821 | 160 |
"nat_rec c d = wfrec pred_nat (%f. nat_case c (%m. d m (f m)))"; |
2608 | 161 |
by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1); |
162 |
bind_thm("nat_rec_unfold", wf_pred_nat RS |
|
163 |
((result() RS eq_reflection) RS def_wfrec)); |
|
164 |
||
165 |
(*--------------------------------------------------------------------------- |
|
166 |
* Old: |
|
167 |
* bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); |
|
168 |
*---------------------------------------------------------------------------*) |
|
169 |
||
170 |
(** conversion rules **) |
|
171 |
||
5069 | 172 |
Goal "nat_rec c h 0 = c"; |
2608 | 173 |
by (rtac (nat_rec_unfold RS trans) 1); |
4089 | 174 |
by (simp_tac (simpset() addsimps [nat_case_0]) 1); |
2608 | 175 |
qed "nat_rec_0"; |
176 |
||
5069 | 177 |
Goal "nat_rec c h (Suc n) = h n (nat_rec c h n)"; |
2608 | 178 |
by (rtac (nat_rec_unfold RS trans) 1); |
4089 | 179 |
by (simp_tac (simpset() addsimps [nat_case_Suc, pred_nat_def, cut_apply]) 1); |
2608 | 180 |
qed "nat_rec_Suc"; |
181 |
||
182 |
(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) |
|
183 |
val [rew] = goal thy |
|
184 |
"[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c"; |
|
185 |
by (rewtac rew); |
|
186 |
by (rtac nat_rec_0 1); |
|
187 |
qed "def_nat_rec_0"; |
|
188 |
||
189 |
val [rew] = goal thy |
|
190 |
"[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)"; |
|
191 |
by (rewtac rew); |
|
192 |
by (rtac nat_rec_Suc 1); |
|
193 |
qed "def_nat_rec_Suc"; |
|
194 |
||
195 |
fun nat_recs def = |
|
196 |
[standard (def RS def_nat_rec_0), |
|
197 |
standard (def RS def_nat_rec_Suc)]; |
|
198 |
||
199 |
||
200 |
(*** Basic properties of "less than" ***) |
|
201 |
||
3378 | 202 |
(*Used in TFL/post.sml*) |
5069 | 203 |
Goalw [less_def] "(m,n) : pred_nat^+ = (m<n)"; |
3378 | 204 |
by (rtac refl 1); |
205 |
qed "less_eq"; |
|
206 |
||
2608 | 207 |
(** Introduction properties **) |
208 |
||
209 |
val prems = goalw thy [less_def] "[| i<j; j<k |] ==> i<(k::nat)"; |
|
210 |
by (rtac (trans_trancl RS transD) 1); |
|
211 |
by (resolve_tac prems 1); |
|
212 |
by (resolve_tac prems 1); |
|
213 |
qed "less_trans"; |
|
214 |
||
5069 | 215 |
Goalw [less_def, pred_nat_def] "n < Suc(n)"; |
4089 | 216 |
by (simp_tac (simpset() addsimps [r_into_trancl]) 1); |
2608 | 217 |
qed "lessI"; |
218 |
AddIffs [lessI]; |
|
219 |
||
220 |
(* i<j ==> i<Suc(j) *) |
|
221 |
bind_thm("less_SucI", lessI RSN (2, less_trans)); |
|
222 |
Addsimps [less_SucI]; |
|
223 |
||
5069 | 224 |
Goal "0 < Suc(n)"; |
2608 | 225 |
by (nat_ind_tac "n" 1); |
226 |
by (rtac lessI 1); |
|
227 |
by (etac less_trans 1); |
|
228 |
by (rtac lessI 1); |
|
229 |
qed "zero_less_Suc"; |
|
230 |
AddIffs [zero_less_Suc]; |
|
231 |
||
232 |
(** Elimination properties **) |
|
233 |
||
234 |
val prems = goalw thy [less_def] "n<m ==> ~ m<(n::nat)"; |
|
4089 | 235 |
by (blast_tac (claset() addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1); |
2608 | 236 |
qed "less_not_sym"; |
237 |
||
238 |
(* [| n<m; m<n |] ==> R *) |
|
239 |
bind_thm ("less_asym", (less_not_sym RS notE)); |
|
240 |
||
5069 | 241 |
Goalw [less_def] "~ n<(n::nat)"; |
2608 | 242 |
by (rtac notI 1); |
243 |
by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1); |
|
244 |
qed "less_not_refl"; |
|
245 |
||
246 |
(* n<n ==> R *) |
|
247 |
bind_thm ("less_irrefl", (less_not_refl RS notE)); |
|
248 |
||
5069 | 249 |
Goal "!!m. n<m ==> m ~= (n::nat)"; |
4089 | 250 |
by (blast_tac (claset() addSEs [less_irrefl]) 1); |
2608 | 251 |
qed "less_not_refl2"; |
252 |
||
253 |
||
3236 | 254 |
val major::prems = goalw thy [less_def, pred_nat_def] |
2608 | 255 |
"[| i<k; k=Suc(i) ==> P; !!j. [| i<j; k=Suc(j) |] ==> P \ |
256 |
\ |] ==> P"; |
|
257 |
by (rtac (major RS tranclE) 1); |
|
3236 | 258 |
by (ALLGOALS Full_simp_tac); |
2608 | 259 |
by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE' |
3236 | 260 |
eresolve_tac (prems@[asm_rl, Pair_inject]))); |
2608 | 261 |
qed "lessE"; |
262 |
||
5069 | 263 |
Goal "~ n<0"; |
2608 | 264 |
by (rtac notI 1); |
265 |
by (etac lessE 1); |
|
266 |
by (etac Zero_neq_Suc 1); |
|
267 |
by (etac Zero_neq_Suc 1); |
|
268 |
qed "not_less0"; |
|
269 |
||
270 |
AddIffs [not_less0]; |
|
271 |
||
272 |
(* n<0 ==> R *) |
|
273 |
bind_thm ("less_zeroE", not_less0 RS notE); |
|
274 |
||
275 |
val [major,less,eq] = goal thy |
|
276 |
"[| m < Suc(n); m<n ==> P; m=n ==> P |] ==> P"; |
|
277 |
by (rtac (major RS lessE) 1); |
|
278 |
by (rtac eq 1); |
|
2891 | 279 |
by (Blast_tac 1); |
2608 | 280 |
by (rtac less 1); |
2891 | 281 |
by (Blast_tac 1); |
2608 | 282 |
qed "less_SucE"; |
283 |
||
5069 | 284 |
Goal "(m < Suc(n)) = (m < n | m = n)"; |
4089 | 285 |
by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1); |
2608 | 286 |
qed "less_Suc_eq"; |
287 |
||
5069 | 288 |
Goal "(n<1) = (n=0)"; |
4089 | 289 |
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); |
3484
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents:
3457
diff
changeset
|
290 |
qed "less_one"; |
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents:
3457
diff
changeset
|
291 |
AddIffs [less_one]; |
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents:
3457
diff
changeset
|
292 |
|
2608 | 293 |
val prems = goal thy "m<n ==> n ~= 0"; |
294 |
by (res_inst_tac [("n","n")] natE 1); |
|
295 |
by (cut_facts_tac prems 1); |
|
296 |
by (ALLGOALS Asm_full_simp_tac); |
|
297 |
qed "gr_implies_not0"; |
|
4356 | 298 |
|
5069 | 299 |
Goal "(n ~= 0) = (0 < n)"; |
4423 | 300 |
by (rtac natE 1); |
4737 | 301 |
by (Blast_tac 1); |
302 |
by (Blast_tac 1); |
|
4356 | 303 |
qed "neq0_conv"; |
4614 | 304 |
AddIffs [neq0_conv]; |
305 |
||
4635 | 306 |
(*This theorem is useful with blast_tac: (n=0 ==> False) ==> 0<n *) |
307 |
bind_thm ("gr0I", [neq0_conv, notI] MRS iffD1); |
|
2608 | 308 |
|
5069 | 309 |
Goal "(~(0 < n)) = (n=0)"; |
4423 | 310 |
by (rtac iffI 1); |
311 |
by (etac swap 1); |
|
312 |
by (ALLGOALS Asm_full_simp_tac); |
|
4356 | 313 |
qed "not_gr0"; |
314 |
Addsimps [not_gr0]; |
|
315 |
||
5069 | 316 |
Goal "!!m. m<n ==> 0 < n"; |
4356 | 317 |
by (dtac gr_implies_not0 1); |
318 |
by (Asm_full_simp_tac 1); |
|
319 |
qed "gr_implies_gr0"; |
|
320 |
Addsimps [gr_implies_gr0]; |
|
321 |
||
2608 | 322 |
|
5069 | 323 |
Goal "!!m n. m<n ==> Suc(m) < Suc(n)"; |
2608 | 324 |
by (etac rev_mp 1); |
325 |
by (nat_ind_tac "n" 1); |
|
4089 | 326 |
by (ALLGOALS (fast_tac (claset() addEs [less_trans, lessE]))); |
2608 | 327 |
qed "Suc_mono"; |
328 |
||
329 |
(*"Less than" is a linear ordering*) |
|
5069 | 330 |
Goal "m<n | m=n | n<(m::nat)"; |
2608 | 331 |
by (nat_ind_tac "m" 1); |
332 |
by (nat_ind_tac "n" 1); |
|
333 |
by (rtac (refl RS disjI1 RS disjI2) 1); |
|
334 |
by (rtac (zero_less_Suc RS disjI1) 1); |
|
4089 | 335 |
by (blast_tac (claset() addIs [Suc_mono, less_SucI] addEs [lessE]) 1); |
2608 | 336 |
qed "less_linear"; |
337 |
||
5069 | 338 |
Goal "!!m::nat. (m ~= n) = (m<n | n<m)"; |
4737 | 339 |
by (cut_facts_tac [less_linear] 1); |
340 |
by (blast_tac (claset() addSEs [less_irrefl]) 1); |
|
341 |
qed "nat_neq_iff"; |
|
342 |
||
2608 | 343 |
qed_goal "nat_less_cases" thy |
344 |
"[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m" |
|
2935 | 345 |
( fn [major,eqCase,lessCase] => |
2608 | 346 |
[ |
2935 | 347 |
(rtac (less_linear RS disjE) 1), |
2608 | 348 |
(etac disjE 2), |
2935 | 349 |
(etac lessCase 1), |
350 |
(etac (sym RS eqCase) 1), |
|
351 |
(etac major 1) |
|
2608 | 352 |
]); |
353 |
||
4745 | 354 |
|
355 |
(** Inductive (?) properties **) |
|
356 |
||
5069 | 357 |
Goal "!!m. [| m<n; Suc m ~= n |] ==> Suc(m) < n"; |
4745 | 358 |
by (full_simp_tac (simpset() addsimps [nat_neq_iff]) 1); |
359 |
by (blast_tac (claset() addSEs [less_irrefl, less_SucE] addEs [less_asym]) 1); |
|
360 |
qed "Suc_lessI"; |
|
361 |
||
362 |
val [prem] = goal thy "Suc(m) < n ==> m<n"; |
|
363 |
by (rtac (prem RS rev_mp) 1); |
|
364 |
by (nat_ind_tac "n" 1); |
|
365 |
by (ALLGOALS (fast_tac (claset() addSIs [lessI RS less_SucI] |
|
366 |
addEs [less_trans, lessE]))); |
|
367 |
qed "Suc_lessD"; |
|
368 |
||
369 |
val [major,minor] = goal thy |
|
370 |
"[| Suc(i)<k; !!j. [| i<j; k=Suc(j) |] ==> P \ |
|
371 |
\ |] ==> P"; |
|
372 |
by (rtac (major RS lessE) 1); |
|
373 |
by (etac (lessI RS minor) 1); |
|
374 |
by (etac (Suc_lessD RS minor) 1); |
|
375 |
by (assume_tac 1); |
|
376 |
qed "Suc_lessE"; |
|
377 |
||
5069 | 378 |
Goal "!!m n. Suc(m) < Suc(n) ==> m<n"; |
4745 | 379 |
by (blast_tac (claset() addEs [lessE, make_elim Suc_lessD]) 1); |
380 |
qed "Suc_less_SucD"; |
|
381 |
||
382 |
||
5069 | 383 |
Goal "(Suc(m) < Suc(n)) = (m<n)"; |
4745 | 384 |
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]); |
385 |
qed "Suc_less_eq"; |
|
386 |
Addsimps [Suc_less_eq]; |
|
387 |
||
5069 | 388 |
Goal "~(Suc(n) < n)"; |
4745 | 389 |
by (blast_tac (claset() addEs [Suc_lessD RS less_irrefl]) 1); |
390 |
qed "not_Suc_n_less_n"; |
|
391 |
Addsimps [not_Suc_n_less_n]; |
|
392 |
||
5069 | 393 |
Goal "!!i. i<j ==> j<k --> Suc i < k"; |
4745 | 394 |
by (nat_ind_tac "k" 1); |
395 |
by (ALLGOALS (asm_simp_tac (simpset()))); |
|
396 |
by (asm_simp_tac (simpset() addsimps [less_Suc_eq]) 1); |
|
397 |
by (blast_tac (claset() addDs [Suc_lessD]) 1); |
|
398 |
qed_spec_mp "less_trans_Suc"; |
|
399 |
||
2608 | 400 |
(*Can be used with less_Suc_eq to get n=m | n<m *) |
5069 | 401 |
Goal "(~ m < n) = (n < Suc(m))"; |
2608 | 402 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
403 |
by (ALLGOALS Asm_simp_tac); |
|
404 |
qed "not_less_eq"; |
|
405 |
||
406 |
(*Complete induction, aka course-of-values induction*) |
|
407 |
val prems = goalw thy [less_def] |
|
408 |
"[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |] ==> P(n)"; |
|
409 |
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); |
|
410 |
by (eresolve_tac prems 1); |
|
411 |
qed "less_induct"; |
|
412 |
||
413 |
qed_goal "nat_induct2" thy |
|
414 |
"[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [ |
|
3023 | 415 |
cut_facts_tac prems 1, |
416 |
rtac less_induct 1, |
|
417 |
res_inst_tac [("n","n")] natE 1, |
|
418 |
hyp_subst_tac 1, |
|
419 |
atac 1, |
|
420 |
hyp_subst_tac 1, |
|
421 |
res_inst_tac [("n","x")] natE 1, |
|
422 |
hyp_subst_tac 1, |
|
423 |
atac 1, |
|
424 |
hyp_subst_tac 1, |
|
425 |
resolve_tac prems 1, |
|
426 |
dtac spec 1, |
|
427 |
etac mp 1, |
|
428 |
rtac (lessI RS less_trans) 1, |
|
429 |
rtac (lessI RS Suc_mono) 1]); |
|
2608 | 430 |
|
431 |
(*** Properties of <= ***) |
|
432 |
||
5069 | 433 |
Goalw [le_def] "(m <= n) = (m < Suc n)"; |
2608 | 434 |
by (rtac not_less_eq 1); |
435 |
qed "le_eq_less_Suc"; |
|
436 |
||
3343 | 437 |
(* m<=n ==> m < Suc n *) |
438 |
bind_thm ("le_imp_less_Suc", le_eq_less_Suc RS iffD1); |
|
439 |
||
5069 | 440 |
Goalw [le_def] "0 <= n"; |
2608 | 441 |
by (rtac not_less0 1); |
442 |
qed "le0"; |
|
443 |
||
5069 | 444 |
Goalw [le_def] "~ Suc n <= n"; |
2608 | 445 |
by (Simp_tac 1); |
446 |
qed "Suc_n_not_le_n"; |
|
447 |
||
5069 | 448 |
Goalw [le_def] "(i <= 0) = (i = 0)"; |
2608 | 449 |
by (nat_ind_tac "i" 1); |
450 |
by (ALLGOALS Asm_simp_tac); |
|
451 |
qed "le_0_eq"; |
|
4614 | 452 |
AddIffs [le_0_eq]; |
2608 | 453 |
|
454 |
Addsimps [(*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq, |
|
455 |
Suc_n_not_le_n, |
|
456 |
n_not_Suc_n, Suc_n_not_n, |
|
457 |
nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; |
|
458 |
||
5069 | 459 |
Goal "!!m. (m <= Suc(n)) = (m<=n | m = Suc n)"; |
4089 | 460 |
by (simp_tac (simpset() addsimps [le_eq_less_Suc]) 1); |
461 |
by (blast_tac (claset() addSEs [less_SucE] addIs [less_SucI]) 1); |
|
3355 | 462 |
qed "le_Suc_eq"; |
463 |
||
4614 | 464 |
(* [| m <= Suc n; m <= n ==> R; m = Suc n ==> R |] ==> R *) |
465 |
bind_thm ("le_SucE", le_Suc_eq RS iffD1 RS disjE); |
|
466 |
||
2608 | 467 |
(* |
5069 | 468 |
Goal "(Suc m < n | Suc m = n) = (m < n)"; |
2608 | 469 |
by (stac (less_Suc_eq RS sym) 1); |
470 |
by (rtac Suc_less_eq 1); |
|
471 |
qed "Suc_le_eq"; |
|
472 |
||
473 |
this could make the simpset (with less_Suc_eq added again) more confluent, |
|
474 |
but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...) |
|
475 |
*) |
|
476 |
||
477 |
(*Prevents simplification of f and g: much faster*) |
|
478 |
qed_goal "nat_case_weak_cong" thy |
|
479 |
"m=n ==> nat_case a f m = nat_case a f n" |
|
480 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
481 |
||
482 |
qed_goal "nat_rec_weak_cong" thy |
|
483 |
"m=n ==> nat_rec a f m = nat_rec a f n" |
|
484 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
485 |
||
4830 | 486 |
qed_goal "split_nat_case" thy |
2608 | 487 |
"P(nat_case z s n) = ((n=0 --> P z) & (!m. n = Suc m --> P(s m)))" |
488 |
(fn _ => [nat_ind_tac "n" 1, ALLGOALS Asm_simp_tac]); |
|
489 |
||
490 |
val prems = goalw thy [le_def] "~n<m ==> m<=(n::nat)"; |
|
491 |
by (resolve_tac prems 1); |
|
492 |
qed "leI"; |
|
493 |
||
494 |
val prems = goalw thy [le_def] "m<=n ==> ~ n < (m::nat)"; |
|
495 |
by (resolve_tac prems 1); |
|
496 |
qed "leD"; |
|
497 |
||
498 |
val leE = make_elim leD; |
|
499 |
||
5069 | 500 |
Goal "(~n<m) = (m<=(n::nat))"; |
4089 | 501 |
by (blast_tac (claset() addIs [leI] addEs [leE]) 1); |
2608 | 502 |
qed "not_less_iff_le"; |
503 |
||
5069 | 504 |
Goalw [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; |
2891 | 505 |
by (Blast_tac 1); |
2608 | 506 |
qed "not_leE"; |
507 |
||
5069 | 508 |
Goalw [le_def] "(~n<=m) = (m<(n::nat))"; |
4599 | 509 |
by (Simp_tac 1); |
510 |
qed "not_le_iff_less"; |
|
511 |
||
5069 | 512 |
Goalw [le_def] "!!m. m < n ==> Suc(m) <= n"; |
4089 | 513 |
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); |
514 |
by (blast_tac (claset() addSEs [less_irrefl,less_asym]) 1); |
|
3343 | 515 |
qed "Suc_leI"; (*formerly called lessD*) |
2608 | 516 |
|
5069 | 517 |
Goalw [le_def] "!!m. Suc(m) <= n ==> m <= n"; |
4089 | 518 |
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); |
2608 | 519 |
qed "Suc_leD"; |
520 |
||
521 |
(* stronger version of Suc_leD *) |
|
5069 | 522 |
Goalw [le_def] |
2608 | 523 |
"!!m. Suc m <= n ==> m < n"; |
4089 | 524 |
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); |
2608 | 525 |
by (cut_facts_tac [less_linear] 1); |
2891 | 526 |
by (Blast_tac 1); |
2608 | 527 |
qed "Suc_le_lessD"; |
528 |
||
5069 | 529 |
Goal "(Suc m <= n) = (m < n)"; |
4089 | 530 |
by (blast_tac (claset() addIs [Suc_leI, Suc_le_lessD]) 1); |
2608 | 531 |
qed "Suc_le_eq"; |
532 |
||
5069 | 533 |
Goalw [le_def] "!!m. m <= n ==> m <= Suc n"; |
4089 | 534 |
by (blast_tac (claset() addDs [Suc_lessD]) 1); |
2608 | 535 |
qed "le_SucI"; |
536 |
Addsimps[le_SucI]; |
|
537 |
||
538 |
bind_thm ("le_Suc", not_Suc_n_less_n RS leI); |
|
539 |
||
5069 | 540 |
Goalw [le_def] "!!m. m < n ==> m <= (n::nat)"; |
4089 | 541 |
by (blast_tac (claset() addEs [less_asym]) 1); |
2608 | 542 |
qed "less_imp_le"; |
543 |
||
3343 | 544 |
(** Equivalence of m<=n and m<n | m=n **) |
545 |
||
5069 | 546 |
Goalw [le_def] "!!m. m <= n ==> m < n | m=(n::nat)"; |
2608 | 547 |
by (cut_facts_tac [less_linear] 1); |
4089 | 548 |
by (blast_tac (claset() addEs [less_irrefl,less_asym]) 1); |
2608 | 549 |
qed "le_imp_less_or_eq"; |
550 |
||
5069 | 551 |
Goalw [le_def] "!!m. m<n | m=n ==> m <=(n::nat)"; |
2608 | 552 |
by (cut_facts_tac [less_linear] 1); |
4089 | 553 |
by (blast_tac (claset() addSEs [less_irrefl] addEs [less_asym]) 1); |
2608 | 554 |
qed "less_or_eq_imp_le"; |
555 |
||
5069 | 556 |
Goal "(m <= (n::nat)) = (m < n | m=n)"; |
2608 | 557 |
by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); |
558 |
qed "le_eq_less_or_eq"; |
|
559 |
||
4635 | 560 |
(*Useful with Blast_tac. m=n ==> m<=n *) |
561 |
bind_thm ("eq_imp_le", disjI2 RS less_or_eq_imp_le); |
|
562 |
||
5069 | 563 |
Goal "n <= (n::nat)"; |
4089 | 564 |
by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); |
2608 | 565 |
qed "le_refl"; |
566 |
||
5069 | 567 |
Goal "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; |
4468 | 568 |
by (blast_tac (claset() addSDs [le_imp_less_or_eq] |
569 |
addIs [less_trans]) 1); |
|
2608 | 570 |
qed "le_less_trans"; |
571 |
||
5069 | 572 |
Goal "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; |
4468 | 573 |
by (blast_tac (claset() addSDs [le_imp_less_or_eq] |
574 |
addIs [less_trans]) 1); |
|
2608 | 575 |
qed "less_le_trans"; |
576 |
||
5069 | 577 |
Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; |
4468 | 578 |
by (blast_tac (claset() addSDs [le_imp_less_or_eq] |
579 |
addIs [less_or_eq_imp_le, less_trans]) 1); |
|
2608 | 580 |
qed "le_trans"; |
581 |
||
5069 | 582 |
Goal "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; |
4468 | 583 |
(*order_less_irrefl could make this proof fail*) |
584 |
by (blast_tac (claset() addSDs [le_imp_less_or_eq] |
|
585 |
addSEs [less_irrefl] addEs [less_asym]) 1); |
|
2608 | 586 |
qed "le_anti_sym"; |
587 |
||
5069 | 588 |
Goal "(Suc(n) <= Suc(m)) = (n <= m)"; |
4089 | 589 |
by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); |
2608 | 590 |
qed "Suc_le_mono"; |
591 |
||
592 |
AddIffs [Suc_le_mono]; |
|
593 |
||
594 |
(* Axiom 'order_le_less' of class 'order': *) |
|
5069 | 595 |
Goal "(m::nat) < n = (m <= n & m ~= n)"; |
4737 | 596 |
by (simp_tac (simpset() addsimps [le_def, nat_neq_iff]) 1); |
597 |
by (blast_tac (claset() addSEs [less_asym]) 1); |
|
2608 | 598 |
qed "nat_less_le"; |
599 |
||
4640 | 600 |
(* Axiom 'linorder_linear' of class 'linorder': *) |
5069 | 601 |
Goal "(m::nat) <= n | n <= m"; |
4640 | 602 |
by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); |
603 |
by (cut_facts_tac [less_linear] 1); |
|
604 |
by(Blast_tac 1); |
|
605 |
qed "nat_le_linear"; |
|
606 |
||
607 |
||
608 |
(** max |
|
4599 | 609 |
|
5069 | 610 |
Goalw [max_def] "!!z::nat. (z <= max x y) = (z <= x | z <= y)"; |
4686 | 611 |
by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); |
4599 | 612 |
by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); |
613 |
qed "le_max_iff_disj"; |
|
614 |
||
5069 | 615 |
Goalw [max_def] "!!z::nat. (max x y <= z) = (x <= z & y <= z)"; |
4686 | 616 |
by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); |
4599 | 617 |
by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); |
618 |
qed "max_le_iff_conj"; |
|
619 |
||
620 |
||
621 |
(** min **) |
|
622 |
||
5069 | 623 |
Goalw [min_def] "!!z::nat. (z <= min x y) = (z <= x & z <= y)"; |
4686 | 624 |
by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); |
4599 | 625 |
by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); |
626 |
qed "le_min_iff_conj"; |
|
627 |
||
5069 | 628 |
Goalw [min_def] "!!z::nat. (min x y <= z) = (x <= z | y <= z)"; |
4686 | 629 |
by (simp_tac (simpset() addsimps [not_le_iff_less] addsplits) 1); |
4599 | 630 |
by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); |
631 |
qed "min_le_iff_disj"; |
|
4640 | 632 |
**) |
4599 | 633 |
|
2608 | 634 |
(** LEAST -- the least number operator **) |
635 |
||
5069 | 636 |
Goal "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)"; |
4089 | 637 |
by (blast_tac (claset() addIs [leI] addEs [leE]) 1); |
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
638 |
val lemma = result(); |
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
639 |
|
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
640 |
(* This is an old def of Least for nat, which is derived for compatibility *) |
5069 | 641 |
Goalw [Least_def] |
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
642 |
"(LEAST n::nat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))"; |
4089 | 643 |
by (simp_tac (simpset() addsimps [lemma]) 1); |
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
644 |
qed "Least_nat_def"; |
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
645 |
|
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
646 |
val [prem1,prem2] = goalw thy [Least_nat_def] |
3842 | 647 |
"[| P(k::nat); !!x. x<k ==> ~P(x) |] ==> (LEAST x. P(x)) = k"; |
2608 | 648 |
by (rtac select_equality 1); |
4089 | 649 |
by (blast_tac (claset() addSIs [prem1,prem2]) 1); |
2608 | 650 |
by (cut_facts_tac [less_linear] 1); |
4089 | 651 |
by (blast_tac (claset() addSIs [prem1] addSDs [prem2]) 1); |
2608 | 652 |
qed "Least_equality"; |
653 |
||
3842 | 654 |
val [prem] = goal thy "P(k::nat) ==> P(LEAST x. P(x))"; |
2608 | 655 |
by (rtac (prem RS rev_mp) 1); |
656 |
by (res_inst_tac [("n","k")] less_induct 1); |
|
657 |
by (rtac impI 1); |
|
658 |
by (rtac classical 1); |
|
659 |
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1); |
|
660 |
by (assume_tac 1); |
|
661 |
by (assume_tac 2); |
|
2891 | 662 |
by (Blast_tac 1); |
2608 | 663 |
qed "LeastI"; |
664 |
||
665 |
(*Proof is almost identical to the one above!*) |
|
3842 | 666 |
val [prem] = goal thy "P(k::nat) ==> (LEAST x. P(x)) <= k"; |
2608 | 667 |
by (rtac (prem RS rev_mp) 1); |
668 |
by (res_inst_tac [("n","k")] less_induct 1); |
|
669 |
by (rtac impI 1); |
|
670 |
by (rtac classical 1); |
|
671 |
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1); |
|
672 |
by (assume_tac 1); |
|
673 |
by (rtac le_refl 2); |
|
4089 | 674 |
by (blast_tac (claset() addIs [less_imp_le,le_trans]) 1); |
2608 | 675 |
qed "Least_le"; |
676 |
||
3842 | 677 |
val [prem] = goal thy "k < (LEAST x. P(x)) ==> ~P(k::nat)"; |
2608 | 678 |
by (rtac notI 1); |
679 |
by (etac (rewrite_rule [le_def] Least_le RS notE) 1); |
|
680 |
by (rtac prem 1); |
|
681 |
qed "not_less_Least"; |
|
682 |
||
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
683 |
qed_goalw "Least_Suc" thy [Least_nat_def] |
2608 | 684 |
"!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))" |
685 |
(fn _ => [ |
|
686 |
rtac select_equality 1, |
|
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
687 |
fold_goals_tac [Least_nat_def], |
4089 | 688 |
safe_tac (claset() addSEs [LeastI]), |
2608 | 689 |
rename_tac "j" 1, |
690 |
res_inst_tac [("n","j")] natE 1, |
|
2891 | 691 |
Blast_tac 1, |
4089 | 692 |
blast_tac (claset() addDs [Suc_less_SucD, not_less_Least]) 1, |
2608 | 693 |
rename_tac "k n" 1, |
694 |
res_inst_tac [("n","k")] natE 1, |
|
2891 | 695 |
Blast_tac 1, |
2608 | 696 |
hyp_subst_tac 1, |
3143
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents:
3085
diff
changeset
|
697 |
rewtac Least_nat_def, |
2608 | 698 |
rtac (select_equality RS arg_cong RS sym) 1, |
4153 | 699 |
Safe_tac, |
2608 | 700 |
dtac Suc_mono 1, |
2891 | 701 |
Blast_tac 1, |
2608 | 702 |
cut_facts_tac [less_linear] 1, |
4153 | 703 |
Safe_tac, |
2608 | 704 |
atac 2, |
2891 | 705 |
Blast_tac 2, |
2608 | 706 |
dtac Suc_mono 1, |
2891 | 707 |
Blast_tac 1]); |
2608 | 708 |
|
709 |
||
710 |
(*** Instantiation of transitivity prover ***) |
|
711 |
||
712 |
structure Less_Arith = |
|
713 |
struct |
|
714 |
val nat_leI = leI; |
|
715 |
val nat_leD = leD; |
|
716 |
val lessI = lessI; |
|
717 |
val zero_less_Suc = zero_less_Suc; |
|
718 |
val less_reflE = less_irrefl; |
|
719 |
val less_zeroE = less_zeroE; |
|
720 |
val less_incr = Suc_mono; |
|
721 |
val less_decr = Suc_less_SucD; |
|
722 |
val less_incr_rhs = Suc_mono RS Suc_lessD; |
|
723 |
val less_decr_lhs = Suc_lessD; |
|
724 |
val less_trans_Suc = less_trans_Suc; |
|
3343 | 725 |
val leI = Suc_leI RS (Suc_le_mono RS iffD1); |
2608 | 726 |
val not_lessI = leI RS leD |
727 |
val not_leI = prove_goal thy "!!m::nat. n < m ==> ~ m <= n" |
|
728 |
(fn _ => [etac swap2 1, etac leD 1]); |
|
729 |
val eqI = prove_goal thy "!!m. [| m < Suc n; n < Suc m |] ==> m=n" |
|
730 |
(fn _ => [etac less_SucE 1, |
|
4089 | 731 |
blast_tac (claset() addSDs [Suc_less_SucD] addSEs [less_irrefl] |
2891 | 732 |
addDs [less_trans_Suc]) 1, |
2935 | 733 |
assume_tac 1]); |
2608 | 734 |
val leD = le_eq_less_Suc RS iffD1; |
735 |
val not_lessD = nat_leI RS leD; |
|
736 |
val not_leD = not_leE |
|
737 |
val eqD1 = prove_goal thy "!!n. m = n ==> m < Suc n" |
|
738 |
(fn _ => [etac subst 1, rtac lessI 1]); |
|
739 |
val eqD2 = sym RS eqD1; |
|
740 |
||
741 |
fun is_zero(t) = t = Const("0",Type("nat",[])); |
|
742 |
||
743 |
fun nnb T = T = Type("fun",[Type("nat",[]), |
|
744 |
Type("fun",[Type("nat",[]), |
|
745 |
Type("bool",[])])]) |
|
746 |
||
747 |
fun decomp_Suc(Const("Suc",_)$t) = let val (a,i) = decomp_Suc t in (a,i+1) end |
|
748 |
| decomp_Suc t = (t,0); |
|
749 |
||
750 |
fun decomp2(rel,T,lhs,rhs) = |
|
751 |
if not(nnb T) then None else |
|
752 |
let val (x,i) = decomp_Suc lhs |
|
753 |
val (y,j) = decomp_Suc rhs |
|
754 |
in case rel of |
|
755 |
"op <" => Some(x,i,"<",y,j) |
|
756 |
| "op <=" => Some(x,i,"<=",y,j) |
|
757 |
| "op =" => Some(x,i,"=",y,j) |
|
758 |
| _ => None |
|
759 |
end; |
|
760 |
||
761 |
fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j) |
|
762 |
| negate None = None; |
|
763 |
||
764 |
fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp2(rel,T,lhs,rhs) |
|
2718 | 765 |
| decomp(_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) = |
2608 | 766 |
negate(decomp2(rel,T,lhs,rhs)) |
767 |
| decomp _ = None |
|
768 |
||
769 |
end; |
|
770 |
||
771 |
structure Trans_Tac = Trans_Tac_Fun(Less_Arith); |
|
772 |
||
773 |
open Trans_Tac; |
|
774 |
||
775 |
(*** eliminates ~= in premises, which trans_tac cannot deal with ***) |
|
4737 | 776 |
bind_thm("nat_neqE", nat_neq_iff RS iffD1 RS disjE); |
2680
20fa49e610ca
function nat_add_primrec added to allow primrec definitions over nat
pusch
parents:
2608
diff
changeset
|
777 |
|
20fa49e610ca
function nat_add_primrec added to allow primrec definitions over nat
pusch
parents:
2608
diff
changeset
|
778 |
|
20fa49e610ca
function nat_add_primrec added to allow primrec definitions over nat
pusch
parents:
2608
diff
changeset
|
779 |
(* add function nat_add_primrec *) |
4737 | 780 |
val (_, nat_add_primrec, _, _) = |
781 |
Datatype.add_datatype ([], "nat", |
|
782 |
[("0", [], Mixfix ("0", [], max_pri)), |
|
783 |
("Suc", [dtTyp ([], "nat")], NoSyn)]) |
|
784 |
(Theory.add_name "Arith" HOL.thy); |
|
785 |
||
3768 | 786 |
(*pretend Arith is part of the basic theory to fool package*) |