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