author | lcp |
Thu, 06 Apr 1995 11:21:13 +0200 | |
changeset 239 | e65129244146 |
parent 233 | f02021cf7cec |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/nat |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
|
4 |
Copyright 1991 University of Cambridge |
|
5 |
||
6 |
For nat.thy. Type nat is defined as a set (Nat) over the type ind. |
|
7 |
*) |
|
8 |
||
9 |
open Nat; |
|
10 |
||
11 |
goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))"; |
|
12 |
by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1)); |
|
171 | 13 |
qed "Nat_fun_mono"; |
0 | 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*) |
|
18 |
goal Nat.thy "Zero_Rep: Nat"; |
|
19 |
by (rtac (Nat_unfold RS ssubst) 1); |
|
20 |
by (rtac (singletonI RS UnI1) 1); |
|
171 | 21 |
qed "Zero_RepI"; |
0 | 22 |
|
23 |
val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat"; |
|
24 |
by (rtac (Nat_unfold RS ssubst) 1); |
|
25 |
by (rtac (imageI RS UnI2) 1); |
|
26 |
by (resolve_tac prems 1); |
|
171 | 27 |
qed "Suc_RepI"; |
0 | 28 |
|
29 |
(*** Induction ***) |
|
30 |
||
31 |
val major::prems = goal Nat.thy |
|
32 |
"[| i: Nat; P(Zero_Rep); \ |
|
33 |
\ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)"; |
|
128 | 34 |
by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); |
0 | 35 |
by (fast_tac (set_cs addIs prems) 1); |
171 | 36 |
qed "Nat_induct"; |
0 | 37 |
|
38 |
val prems = goalw Nat.thy [Zero_def,Suc_def] |
|
39 |
"[| P(0); \ |
|
40 |
\ !!k. P(k) ==> P(Suc(k)) |] ==> P(n)"; |
|
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)); |
|
171 | 45 |
qed "nat_induct"; |
0 | 46 |
|
47 |
(*Perform induction on n. *) |
|
48 |
fun nat_ind_tac a i = |
|
49 |
EVERY [res_inst_tac [("n",a)] nat_induct i, |
|
50 |
rename_last_tac a ["1"] (i+1)]; |
|
51 |
||
52 |
(*A special form of induction for reasoning about m<n and m-n*) |
|
53 |
val prems = goal Nat.thy |
|
54 |
"[| !!x. P(x,0); \ |
|
55 |
\ !!y. P(0,Suc(y)); \ |
|
56 |
\ !!x y. [| P(x,y) |] ==> P(Suc(x),Suc(y)) \ |
|
57 |
\ |] ==> P(m,n)"; |
|
58 |
by (res_inst_tac [("x","m")] spec 1); |
|
59 |
by (nat_ind_tac "n" 1); |
|
60 |
by (rtac allI 2); |
|
61 |
by (nat_ind_tac "x" 2); |
|
62 |
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1)); |
|
171 | 63 |
qed "diff_induct"; |
0 | 64 |
|
65 |
(*Case analysis on the natural numbers*) |
|
66 |
val prems = goal Nat.thy |
|
67 |
"[| n=0 ==> P; !!x. n = Suc(x) ==> P |] ==> P"; |
|
68 |
by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1); |
|
69 |
by (fast_tac (HOL_cs addSEs prems) 1); |
|
70 |
by (nat_ind_tac "n" 1); |
|
71 |
by (rtac (refl RS disjI1) 1); |
|
72 |
by (fast_tac HOL_cs 1); |
|
171 | 73 |
qed "natE"; |
0 | 74 |
|
75 |
(*** Isomorphisms: Abs_Nat and Rep_Nat ***) |
|
76 |
||
77 |
(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat), |
|
78 |
since we assume the isomorphism equations will one day be given by Isabelle*) |
|
79 |
||
80 |
goal Nat.thy "inj(Rep_Nat)"; |
|
81 |
by (rtac inj_inverseI 1); |
|
82 |
by (rtac Rep_Nat_inverse 1); |
|
171 | 83 |
qed "inj_Rep_Nat"; |
0 | 84 |
|
85 |
goal Nat.thy "inj_onto(Abs_Nat,Nat)"; |
|
86 |
by (rtac inj_onto_inverseI 1); |
|
87 |
by (etac Abs_Nat_inverse 1); |
|
171 | 88 |
qed "inj_onto_Abs_Nat"; |
0 | 89 |
|
90 |
(*** Distinctness of constructors ***) |
|
91 |
||
5 | 92 |
goalw Nat.thy [Zero_def,Suc_def] "Suc(m) ~= 0"; |
0 | 93 |
by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1); |
94 |
by (rtac Suc_Rep_not_Zero_Rep 1); |
|
95 |
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); |
|
171 | 96 |
qed "Suc_not_Zero"; |
0 | 97 |
|
202 | 98 |
bind_thm ("Zero_not_Suc", (Suc_not_Zero RS not_sym)); |
0 | 99 |
|
202 | 100 |
bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE)); |
0 | 101 |
val Zero_neq_Suc = sym RS Suc_neq_Zero; |
102 |
||
103 |
(** Injectiveness of Suc **) |
|
104 |
||
105 |
goalw Nat.thy [Suc_def] "inj(Suc)"; |
|
106 |
by (rtac injI 1); |
|
107 |
by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1); |
|
108 |
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); |
|
109 |
by (dtac (inj_Suc_Rep RS injD) 1); |
|
110 |
by (etac (inj_Rep_Nat RS injD) 1); |
|
171 | 111 |
qed "inj_Suc"; |
0 | 112 |
|
113 |
val Suc_inject = inj_Suc RS injD;; |
|
114 |
||
115 |
goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)"; |
|
116 |
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); |
|
171 | 117 |
qed "Suc_Suc_eq"; |
0 | 118 |
|
5 | 119 |
goal Nat.thy "n ~= Suc(n)"; |
0 | 120 |
by (nat_ind_tac "n" 1); |
121 |
by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [Zero_not_Suc,Suc_Suc_eq]))); |
|
171 | 122 |
qed "n_not_Suc_n"; |
0 | 123 |
|
45 | 124 |
val Suc_n_not_n = n_not_Suc_n RS not_sym; |
125 |
||
0 | 126 |
(*** nat_case -- the selection operator for nat ***) |
127 |
||
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
128 |
goalw Nat.thy [nat_case_def] "nat_case(a, f, 0) = a"; |
0 | 129 |
by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); |
171 | 130 |
qed "nat_case_0"; |
0 | 131 |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
132 |
goalw Nat.thy [nat_case_def] "nat_case(a, f, Suc(k)) = f(k)"; |
0 | 133 |
by (fast_tac (set_cs addIs [select_equality] |
134 |
addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); |
|
171 | 135 |
qed "nat_case_Suc"; |
0 | 136 |
|
137 |
(** Introduction rules for 'pred_nat' **) |
|
138 |
||
139 |
goalw Nat.thy [pred_nat_def] "<n, Suc(n)> : pred_nat"; |
|
140 |
by (fast_tac set_cs 1); |
|
171 | 141 |
qed "pred_natI"; |
0 | 142 |
|
143 |
val major::prems = goalw Nat.thy [pred_nat_def] |
|
144 |
"[| p : pred_nat; !!x n. [| p = <n, Suc(n)> |] ==> R \ |
|
145 |
\ |] ==> R"; |
|
146 |
by (rtac (major RS CollectE) 1); |
|
147 |
by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1)); |
|
171 | 148 |
qed "pred_natE"; |
0 | 149 |
|
150 |
goalw Nat.thy [wf_def] "wf(pred_nat)"; |
|
151 |
by (strip_tac 1); |
|
152 |
by (nat_ind_tac "x" 1); |
|
153 |
by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, |
|
154 |
make_elim Suc_inject]) 2); |
|
155 |
by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1); |
|
171 | 156 |
qed "wf_pred_nat"; |
0 | 157 |
|
158 |
||
159 |
(*** nat_rec -- by wf recursion on pred_nat ***) |
|
160 |
||
202 | 161 |
bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); |
0 | 162 |
|
163 |
(** conversion rules **) |
|
164 |
||
165 |
goal Nat.thy "nat_rec(0,c,h) = c"; |
|
166 |
by (rtac (nat_rec_unfold RS trans) 1); |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
167 |
by (simp_tac (HOL_ss addsimps [nat_case_0]) 1); |
171 | 168 |
qed "nat_rec_0"; |
0 | 169 |
|
170 |
goal Nat.thy "nat_rec(Suc(n), c, h) = h(n, nat_rec(n,c,h))"; |
|
171 |
by (rtac (nat_rec_unfold RS trans) 1); |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
172 |
by (simp_tac (HOL_ss addsimps [nat_case_Suc, pred_natI, cut_apply]) 1); |
171 | 173 |
qed "nat_rec_Suc"; |
0 | 174 |
|
175 |
(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) |
|
176 |
val [rew] = goal Nat.thy |
|
177 |
"[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(0) = c"; |
|
178 |
by (rewtac rew); |
|
179 |
by (rtac nat_rec_0 1); |
|
171 | 180 |
qed "def_nat_rec_0"; |
0 | 181 |
|
182 |
val [rew] = goal Nat.thy |
|
183 |
"[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(Suc(n)) = h(n,f(n))"; |
|
184 |
by (rewtac rew); |
|
185 |
by (rtac nat_rec_Suc 1); |
|
171 | 186 |
qed "def_nat_rec_Suc"; |
0 | 187 |
|
188 |
fun nat_recs def = |
|
189 |
[standard (def RS def_nat_rec_0), |
|
190 |
standard (def RS def_nat_rec_Suc)]; |
|
191 |
||
192 |
||
193 |
(*** Basic properties of "less than" ***) |
|
194 |
||
195 |
(** Introduction properties **) |
|
196 |
||
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
197 |
val prems = goalw Nat.thy [less_def] "[| i<j; j<k |] ==> i<(k::nat)"; |
0 | 198 |
by (rtac (trans_trancl RS transD) 1); |
199 |
by (resolve_tac prems 1); |
|
200 |
by (resolve_tac prems 1); |
|
171 | 201 |
qed "less_trans"; |
0 | 202 |
|
203 |
goalw Nat.thy [less_def] "n < Suc(n)"; |
|
204 |
by (rtac (pred_natI RS r_into_trancl) 1); |
|
171 | 205 |
qed "lessI"; |
0 | 206 |
|
207 |
(* i<j ==> i<Suc(j) *) |
|
208 |
val less_SucI = lessI RSN (2, less_trans); |
|
209 |
||
210 |
goal Nat.thy "0 < Suc(n)"; |
|
211 |
by (nat_ind_tac "n" 1); |
|
212 |
by (rtac lessI 1); |
|
213 |
by (etac less_trans 1); |
|
214 |
by (rtac lessI 1); |
|
171 | 215 |
qed "zero_less_Suc"; |
0 | 216 |
|
217 |
(** Elimination properties **) |
|
218 |
||
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
219 |
val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)"; |
89 | 220 |
by(fast_tac (HOL_cs addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1); |
171 | 221 |
qed "less_not_sym"; |
0 | 222 |
|
223 |
(* [| n<m; m<n |] ==> R *) |
|
202 | 224 |
bind_thm ("less_asym", (less_not_sym RS notE)); |
0 | 225 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
226 |
goalw Nat.thy [less_def] "~ n<(n::nat)"; |
0 | 227 |
by (rtac notI 1); |
228 |
by (etac (wf_pred_nat RS wf_trancl RS wf_anti_refl) 1); |
|
171 | 229 |
qed "less_not_refl"; |
0 | 230 |
|
231 |
(* n<n ==> R *) |
|
202 | 232 |
bind_thm ("less_anti_refl", (less_not_refl RS notE)); |
0 | 233 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
234 |
goal Nat.thy "!!m. n<m ==> m ~= (n::nat)"; |
45 | 235 |
by(fast_tac (HOL_cs addEs [less_anti_refl]) 1); |
171 | 236 |
qed "less_not_refl2"; |
45 | 237 |
|
0 | 238 |
|
239 |
val major::prems = goalw Nat.thy [less_def] |
|
240 |
"[| i<k; k=Suc(i) ==> P; !!j. [| i<j; k=Suc(j) |] ==> P \ |
|
241 |
\ |] ==> P"; |
|
242 |
by (rtac (major RS tranclE) 1); |
|
239 | 243 |
by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE' |
244 |
eresolve_tac (prems@[pred_natE, Pair_inject]))); |
|
245 |
by (rtac refl 1); |
|
171 | 246 |
qed "lessE"; |
0 | 247 |
|
248 |
goal Nat.thy "~ n<0"; |
|
249 |
by (rtac notI 1); |
|
250 |
by (etac lessE 1); |
|
251 |
by (etac Zero_neq_Suc 1); |
|
252 |
by (etac Zero_neq_Suc 1); |
|
171 | 253 |
qed "not_less0"; |
0 | 254 |
|
255 |
(* n<0 ==> R *) |
|
202 | 256 |
bind_thm ("less_zeroE", (not_less0 RS notE)); |
0 | 257 |
|
258 |
val [major,less,eq] = goal Nat.thy |
|
259 |
"[| m < Suc(n); m<n ==> P; m=n ==> P |] ==> P"; |
|
260 |
by (rtac (major RS lessE) 1); |
|
261 |
by (rtac eq 1); |
|
262 |
by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); |
|
263 |
by (rtac less 1); |
|
264 |
by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); |
|
171 | 265 |
qed "less_SucE"; |
0 | 266 |
|
267 |
goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; |
|
268 |
by (fast_tac (HOL_cs addSIs [lessI] |
|
269 |
addEs [less_trans, less_SucE]) 1); |
|
171 | 270 |
qed "less_Suc_eq"; |
0 | 271 |
|
272 |
||
273 |
(** Inductive (?) properties **) |
|
274 |
||
275 |
val [prem] = goal Nat.thy "Suc(m) < n ==> m<n"; |
|
276 |
by (rtac (prem RS rev_mp) 1); |
|
277 |
by (nat_ind_tac "n" 1); |
|
278 |
by (rtac impI 1); |
|
279 |
by (etac less_zeroE 1); |
|
280 |
by (fast_tac (HOL_cs addSIs [lessI RS less_SucI] |
|
281 |
addSDs [Suc_inject] |
|
282 |
addEs [less_trans, lessE]) 1); |
|
171 | 283 |
qed "Suc_lessD"; |
0 | 284 |
|
285 |
val [major,minor] = goal Nat.thy |
|
286 |
"[| Suc(i)<k; !!j. [| i<j; k=Suc(j) |] ==> P \ |
|
287 |
\ |] ==> P"; |
|
288 |
by (rtac (major RS lessE) 1); |
|
289 |
by (etac (lessI RS minor) 1); |
|
290 |
by (etac (Suc_lessD RS minor) 1); |
|
291 |
by (assume_tac 1); |
|
171 | 292 |
qed "Suc_lessE"; |
0 | 293 |
|
294 |
val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m<n"; |
|
295 |
by (rtac (major RS lessE) 1); |
|
296 |
by (REPEAT (rtac lessI 1 |
|
297 |
ORELSE eresolve_tac [make_elim Suc_inject, ssubst, Suc_lessD] 1)); |
|
171 | 298 |
qed "Suc_less_SucD"; |
0 | 299 |
|
300 |
val prems = goal Nat.thy "m<n ==> Suc(m) < Suc(n)"; |
|
301 |
by (subgoal_tac "m<n --> Suc(m) < Suc(n)" 1); |
|
302 |
by (fast_tac (HOL_cs addIs prems) 1); |
|
303 |
by (nat_ind_tac "n" 1); |
|
304 |
by (rtac impI 1); |
|
305 |
by (etac less_zeroE 1); |
|
306 |
by (fast_tac (HOL_cs addSIs [lessI] |
|
307 |
addSDs [Suc_inject] |
|
308 |
addEs [less_trans, lessE]) 1); |
|
171 | 309 |
qed "Suc_mono"; |
0 | 310 |
|
311 |
goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)"; |
|
312 |
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]); |
|
171 | 313 |
qed "Suc_less_eq"; |
0 | 314 |
|
45 | 315 |
goal Nat.thy "~(Suc(n) < n)"; |
316 |
by(fast_tac (HOL_cs addEs [Suc_lessD RS less_anti_refl]) 1); |
|
171 | 317 |
qed "not_Suc_n_less_n"; |
0 | 318 |
|
319 |
(*"Less than" is a linear ordering*) |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
320 |
goal Nat.thy "m<n | m=n | n<(m::nat)"; |
0 | 321 |
by (nat_ind_tac "m" 1); |
322 |
by (nat_ind_tac "n" 1); |
|
323 |
by (rtac (refl RS disjI1 RS disjI2) 1); |
|
324 |
by (rtac (zero_less_Suc RS disjI1) 1); |
|
325 |
by (fast_tac (HOL_cs addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1); |
|
171 | 326 |
qed "less_linear"; |
0 | 327 |
|
328 |
(*Can be used with less_Suc_eq to get n=m | n<m *) |
|
329 |
goal Nat.thy "(~ m < n) = (n < Suc(m))"; |
|
330 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
331 |
by(ALLGOALS(asm_simp_tac (HOL_ss addsimps |
|
332 |
[not_less0,zero_less_Suc,Suc_less_eq]))); |
|
171 | 333 |
qed "not_less_eq"; |
0 | 334 |
|
335 |
(*Complete induction, aka course-of-values induction*) |
|
336 |
val prems = goalw Nat.thy [less_def] |
|
337 |
"[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |] ==> P(n)"; |
|
338 |
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); |
|
339 |
by (eresolve_tac prems 1); |
|
171 | 340 |
qed "less_induct"; |
0 | 341 |
|
342 |
||
343 |
(*** Properties of <= ***) |
|
344 |
||
345 |
goalw Nat.thy [le_def] "0 <= n"; |
|
346 |
by (rtac not_less0 1); |
|
171 | 347 |
qed "le0"; |
0 | 348 |
|
349 |
val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
350 |
Suc_less_eq, less_Suc_eq, le0, not_Suc_n_less_n, |
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
351 |
Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, |
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
352 |
n_not_Suc_n, Suc_n_not_n, |
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
353 |
nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; |
0 | 354 |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
355 |
val nat_ss0 = sum_ss addsimps nat_simps; |
0 | 356 |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
357 |
(*Prevents simplification of f and g: much faster*) |
179 | 358 |
qed_goal "nat_case_weak_cong" Nat.thy |
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
359 |
"m=n ==> nat_case(a,f,m) = nat_case(a,f,n)" |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
360 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
361 |
|
179 | 362 |
qed_goal "nat_rec_weak_cong" Nat.thy |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
363 |
"m=n ==> nat_rec(m,a,f) = nat_rec(n,a,f)" |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
364 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
365 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
366 |
val prems = goalw Nat.thy [le_def] "~(n<m) ==> m<=(n::nat)"; |
0 | 367 |
by (resolve_tac prems 1); |
171 | 368 |
qed "leI"; |
0 | 369 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
370 |
val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<(m::nat))"; |
0 | 371 |
by (resolve_tac prems 1); |
171 | 372 |
qed "leD"; |
0 | 373 |
|
374 |
val leE = make_elim leD; |
|
375 |
||
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
376 |
goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; |
0 | 377 |
by (fast_tac HOL_cs 1); |
171 | 378 |
qed "not_leE"; |
0 | 379 |
|
380 |
goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n"; |
|
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
381 |
by(simp_tac nat_ss0 1); |
89 | 382 |
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); |
171 | 383 |
qed "lessD"; |
0 | 384 |
|
40 | 385 |
goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n"; |
109
c53c19fb22cb
HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents:
90
diff
changeset
|
386 |
by(asm_full_simp_tac nat_ss0 1); |
40 | 387 |
by(fast_tac HOL_cs 1); |
171 | 388 |
qed "Suc_leD"; |
40 | 389 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
390 |
goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)"; |
89 | 391 |
by (fast_tac (HOL_cs addEs [less_asym]) 1); |
171 | 392 |
qed "less_imp_le"; |
0 | 393 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
394 |
goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)"; |
0 | 395 |
by (cut_facts_tac [less_linear] 1); |
89 | 396 |
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); |
171 | 397 |
qed "le_imp_less_or_eq"; |
0 | 398 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
399 |
goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)"; |
0 | 400 |
by (cut_facts_tac [less_linear] 1); |
89 | 401 |
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); |
0 | 402 |
by (flexflex_tac); |
171 | 403 |
qed "less_or_eq_imp_le"; |
0 | 404 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
405 |
goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)"; |
0 | 406 |
by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); |
171 | 407 |
qed "le_eq_less_or_eq"; |
0 | 408 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
409 |
goal Nat.thy "n <= (n::nat)"; |
0 | 410 |
by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); |
171 | 411 |
qed "le_refl"; |
0 | 412 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
413 |
val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; |
0 | 414 |
by (dtac le_imp_less_or_eq 1); |
415 |
by (fast_tac (HOL_cs addIs [less_trans]) 1); |
|
171 | 416 |
qed "le_less_trans"; |
0 | 417 |
|
218
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
418 |
goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
419 |
by (dtac le_imp_less_or_eq 1); |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
420 |
by (fast_tac (HOL_cs addIs [less_trans]) 1); |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
421 |
qed "less_le_trans"; |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
422 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
423 |
goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; |
0 | 424 |
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, |
425 |
rtac less_or_eq_imp_le, fast_tac (HOL_cs addIs [less_trans])]); |
|
171 | 426 |
qed "le_trans"; |
0 | 427 |
|
90
5c7a69cef18b
added parentheses made necessary by change of constrain's precedence
clasohm
parents:
89
diff
changeset
|
428 |
val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; |
0 | 429 |
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, |
89 | 430 |
fast_tac (HOL_cs addEs [less_anti_refl,less_asym])]); |
171 | 431 |
qed "le_anti_sym"; |
57 | 432 |
|
218
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
433 |
goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)"; |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
434 |
by (simp_tac (nat_ss0 addsimps [le_eq_less_or_eq]) 1); |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
435 |
qed "Suc_le_mono"; |
78c9acf5bc38
Moved succ_leq_mono from IOA/example/Lemmas to Nat.ML as Suc_le_mono.
lcp
parents:
202
diff
changeset
|
436 |
|
233 | 437 |
val nat_ss = nat_ss0 addsimps [le_refl,Suc_le_mono]; |