author | paulson |
Thu, 20 Nov 1997 11:03:26 +0100 | |
changeset 4242 | 97601cf26262 |
parent 4091 | 771b1f6422a8 |
child 4839 | a7322db15065 |
permissions | -rw-r--r-- |
1793 | 1 |
(* Title: ZF/Arith.ML |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
1609 | 6 |
Arithmetic operators and their definitions |
0 | 7 |
|
8 |
Proofs about elementary arithmetic: addition, multiplication, etc. |
|
9 |
||
10 |
Could prove def_rec_0, def_rec_succ... |
|
11 |
*) |
|
12 |
||
13 |
open Arith; |
|
14 |
||
15 |
(*"Difference" is subtraction of natural numbers. |
|
16 |
There are no negative numbers; we have |
|
17 |
m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. |
|
18 |
Also, rec(m, 0, %z w.z) is pred(m). |
|
19 |
*) |
|
20 |
||
21 |
(** rec -- better than nat_rec; the succ case has no type requirement! **) |
|
22 |
||
23 |
val rec_trans = rec_def RS def_transrec RS trans; |
|
24 |
||
25 |
goal Arith.thy "rec(0,a,b) = a"; |
|
26 |
by (rtac rec_trans 1); |
|
27 |
by (rtac nat_case_0 1); |
|
760 | 28 |
qed "rec_0"; |
0 | 29 |
|
30 |
goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))"; |
|
31 |
by (rtac rec_trans 1); |
|
2469 | 32 |
by (Simp_tac 1); |
760 | 33 |
qed "rec_succ"; |
0 | 34 |
|
2469 | 35 |
Addsimps [rec_0, rec_succ]; |
36 |
||
0 | 37 |
val major::prems = goal Arith.thy |
38 |
"[| n: nat; \ |
|
39 |
\ a: C(0); \ |
|
40 |
\ !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) \ |
|
41 |
\ |] ==> rec(n,a,b) : C(n)"; |
|
42 |
by (rtac (major RS nat_induct) 1); |
|
43 |
by (ALLGOALS |
|
4091 | 44 |
(asm_simp_tac (simpset() addsimps prems))); |
760 | 45 |
qed "rec_type"; |
0 | 46 |
|
2469 | 47 |
Addsimps [rec_type, nat_0_le, nat_le_refl]; |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
48 |
val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]; |
0 | 49 |
|
1708 | 50 |
goal Arith.thy "!!k. [| 0<k; k: nat |] ==> EX j: nat. k = succ(j)"; |
51 |
by (etac rev_mp 1); |
|
52 |
by (etac nat_induct 1); |
|
2469 | 53 |
by (Simp_tac 1); |
3016 | 54 |
by (Blast_tac 1); |
1708 | 55 |
val lemma = result(); |
56 |
||
57 |
(* [| 0 < k; k: nat; !!j. [| j: nat; k = succ(j) |] ==> Q |] ==> Q *) |
|
58 |
bind_thm ("zero_lt_natE", lemma RS bexE); |
|
59 |
||
60 |
||
0 | 61 |
(** Addition **) |
62 |
||
760 | 63 |
qed_goalw "add_type" Arith.thy [add_def] |
0 | 64 |
"[| m:nat; n:nat |] ==> m #+ n : nat" |
65 |
(fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); |
|
66 |
||
760 | 67 |
qed_goalw "add_0" Arith.thy [add_def] |
0 | 68 |
"0 #+ n = n" |
69 |
(fn _ => [ (rtac rec_0 1) ]); |
|
70 |
||
760 | 71 |
qed_goalw "add_succ" Arith.thy [add_def] |
0 | 72 |
"succ(m) #+ n = succ(m #+ n)" |
73 |
(fn _=> [ (rtac rec_succ 1) ]); |
|
74 |
||
2469 | 75 |
Addsimps [add_0, add_succ]; |
76 |
||
0 | 77 |
(** Multiplication **) |
78 |
||
760 | 79 |
qed_goalw "mult_type" Arith.thy [mult_def] |
0 | 80 |
"[| m:nat; n:nat |] ==> m #* n : nat" |
81 |
(fn prems=> |
|
82 |
[ (typechk_tac (prems@[add_type]@nat_typechecks@ZF_typechecks)) ]); |
|
83 |
||
760 | 84 |
qed_goalw "mult_0" Arith.thy [mult_def] |
0 | 85 |
"0 #* n = 0" |
86 |
(fn _ => [ (rtac rec_0 1) ]); |
|
87 |
||
760 | 88 |
qed_goalw "mult_succ" Arith.thy [mult_def] |
0 | 89 |
"succ(m) #* n = n #+ (m #* n)" |
90 |
(fn _ => [ (rtac rec_succ 1) ]); |
|
91 |
||
2469 | 92 |
Addsimps [mult_0, mult_succ]; |
93 |
||
0 | 94 |
(** Difference **) |
95 |
||
760 | 96 |
qed_goalw "diff_type" Arith.thy [diff_def] |
0 | 97 |
"[| m:nat; n:nat |] ==> m #- n : nat" |
98 |
(fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); |
|
99 |
||
760 | 100 |
qed_goalw "diff_0" Arith.thy [diff_def] |
0 | 101 |
"m #- 0 = m" |
102 |
(fn _ => [ (rtac rec_0 1) ]); |
|
103 |
||
760 | 104 |
qed_goalw "diff_0_eq_0" Arith.thy [diff_def] |
0 | 105 |
"n:nat ==> 0 #- n = 0" |
106 |
(fn [prem]=> |
|
107 |
[ (rtac (prem RS nat_induct) 1), |
|
2469 | 108 |
(ALLGOALS (Asm_simp_tac)) ]); |
0 | 109 |
|
110 |
(*Must simplify BEFORE the induction!! (Else we get a critical pair) |
|
111 |
succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *) |
|
760 | 112 |
qed_goalw "diff_succ_succ" Arith.thy [diff_def] |
0 | 113 |
"[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n" |
114 |
(fn prems=> |
|
4091 | 115 |
[ (asm_simp_tac (simpset() addsimps prems) 1), |
0 | 116 |
(nat_ind_tac "n" prems 1), |
4091 | 117 |
(ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); |
2469 | 118 |
|
119 |
Addsimps [diff_0, diff_0_eq_0, diff_succ_succ]; |
|
0 | 120 |
|
121 |
val prems = goal Arith.thy |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
122 |
"[| m:nat; n:nat |] ==> m #- n le m"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
123 |
by (rtac (prems MRS diff_induct) 1); |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
124 |
by (etac leE 3); |
0 | 125 |
by (ALLGOALS |
4091 | 126 |
(asm_simp_tac (simpset() addsimps (prems @ [le_iff, nat_into_Ord])))); |
760 | 127 |
qed "diff_le_self"; |
0 | 128 |
|
129 |
(*** Simplification over add, mult, diff ***) |
|
130 |
||
131 |
val arith_typechecks = [add_type, mult_type, diff_type]; |
|
2469 | 132 |
Addsimps arith_typechecks; |
0 | 133 |
|
134 |
||
135 |
(*** Addition ***) |
|
136 |
||
137 |
(*Associative law for addition*) |
|
760 | 138 |
qed_goal "add_assoc" Arith.thy |
0 | 139 |
"m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)" |
140 |
(fn prems=> |
|
141 |
[ (nat_ind_tac "m" prems 1), |
|
4091 | 142 |
(ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); |
0 | 143 |
|
144 |
(*The following two lemmas are used for add_commute and sometimes |
|
145 |
elsewhere, since they are safe for rewriting.*) |
|
760 | 146 |
qed_goal "add_0_right" Arith.thy |
0 | 147 |
"m:nat ==> m #+ 0 = m" |
148 |
(fn prems=> |
|
149 |
[ (nat_ind_tac "m" prems 1), |
|
4091 | 150 |
(ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); |
0 | 151 |
|
760 | 152 |
qed_goal "add_succ_right" Arith.thy |
0 | 153 |
"m:nat ==> m #+ succ(n) = succ(m #+ n)" |
154 |
(fn prems=> |
|
155 |
[ (nat_ind_tac "m" prems 1), |
|
4091 | 156 |
(ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); |
2469 | 157 |
|
158 |
Addsimps [add_0_right, add_succ_right]; |
|
0 | 159 |
|
160 |
(*Commutative law for addition*) |
|
760 | 161 |
qed_goal "add_commute" Arith.thy |
435 | 162 |
"!!m n. [| m:nat; n:nat |] ==> m #+ n = n #+ m" |
163 |
(fn _ => |
|
164 |
[ (nat_ind_tac "n" [] 1), |
|
2469 | 165 |
(ALLGOALS Asm_simp_tac) ]); |
435 | 166 |
|
437 | 167 |
(*for a/c rewriting*) |
760 | 168 |
qed_goal "add_left_commute" Arith.thy |
437 | 169 |
"!!m n k. [| m:nat; n:nat |] ==> m#+(n#+k)=n#+(m#+k)" |
4091 | 170 |
(fn _ => [asm_simp_tac(simpset() addsimps [add_assoc RS sym, add_commute]) 1]); |
435 | 171 |
|
172 |
(*Addition is an AC-operator*) |
|
173 |
val add_ac = [add_assoc, add_commute, add_left_commute]; |
|
0 | 174 |
|
175 |
(*Cancellation law on the left*) |
|
437 | 176 |
val [eqn,knat] = goal Arith.thy |
177 |
"[| k #+ m = k #+ n; k:nat |] ==> m=n"; |
|
0 | 178 |
by (rtac (eqn RS rev_mp) 1); |
179 |
by (nat_ind_tac "k" [knat] 1); |
|
2469 | 180 |
by (ALLGOALS (Simp_tac)); |
760 | 181 |
qed "add_left_cancel"; |
0 | 182 |
|
183 |
(*** Multiplication ***) |
|
184 |
||
185 |
(*right annihilation in product*) |
|
760 | 186 |
qed_goal "mult_0_right" Arith.thy |
435 | 187 |
"!!m. m:nat ==> m #* 0 = 0" |
188 |
(fn _=> |
|
189 |
[ (nat_ind_tac "m" [] 1), |
|
2469 | 190 |
(ALLGOALS (Asm_simp_tac)) ]); |
0 | 191 |
|
192 |
(*right successor law for multiplication*) |
|
760 | 193 |
qed_goal "mult_succ_right" Arith.thy |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
194 |
"!!m n. [| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)" |
435 | 195 |
(fn _ => |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
196 |
[ (nat_ind_tac "m" [] 1), |
4091 | 197 |
(ALLGOALS (asm_simp_tac (simpset() addsimps add_ac))) ]); |
2469 | 198 |
|
199 |
Addsimps [mult_0_right, mult_succ_right]; |
|
0 | 200 |
|
1793 | 201 |
goal Arith.thy "!!n. n:nat ==> 1 #* n = n"; |
2469 | 202 |
by (Asm_simp_tac 1); |
1793 | 203 |
qed "mult_1"; |
204 |
||
205 |
goal Arith.thy "!!n. n:nat ==> n #* 1 = n"; |
|
2469 | 206 |
by (Asm_simp_tac 1); |
1793 | 207 |
qed "mult_1_right"; |
208 |
||
0 | 209 |
(*Commutative law for multiplication*) |
760 | 210 |
qed_goal "mult_commute" Arith.thy |
2469 | 211 |
"!!m n. [| m:nat; n:nat |] ==> m #* n = n #* m" |
212 |
(fn _=> |
|
213 |
[ (nat_ind_tac "m" [] 1), |
|
214 |
(ALLGOALS Asm_simp_tac) ]); |
|
0 | 215 |
|
216 |
(*addition distributes over multiplication*) |
|
760 | 217 |
qed_goal "add_mult_distrib" Arith.thy |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
218 |
"!!m n. [| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)" |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
219 |
(fn _=> |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
220 |
[ (etac nat_induct 1), |
4091 | 221 |
(ALLGOALS (asm_simp_tac (simpset() addsimps [add_assoc RS sym]))) ]); |
0 | 222 |
|
223 |
(*Distributive law on the left; requires an extra typing premise*) |
|
760 | 224 |
qed_goal "add_mult_distrib_left" Arith.thy |
435 | 225 |
"!!m. [| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)" |
0 | 226 |
(fn prems=> |
435 | 227 |
[ (nat_ind_tac "m" [] 1), |
2469 | 228 |
(Asm_simp_tac 1), |
4091 | 229 |
(asm_simp_tac (simpset() addsimps add_ac) 1) ]); |
0 | 230 |
|
231 |
(*Associative law for multiplication*) |
|
760 | 232 |
qed_goal "mult_assoc" Arith.thy |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
233 |
"!!m n k. [| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)" |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
234 |
(fn _=> |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
235 |
[ (etac nat_induct 1), |
4091 | 236 |
(ALLGOALS (asm_simp_tac (simpset() addsimps [add_mult_distrib]))) ]); |
0 | 237 |
|
437 | 238 |
(*for a/c rewriting*) |
760 | 239 |
qed_goal "mult_left_commute" Arith.thy |
437 | 240 |
"!!m n k. [| m:nat; n:nat; k:nat |] ==> m #* (n #* k) = n #* (m #* k)" |
241 |
(fn _ => [rtac (mult_commute RS trans) 1, |
|
242 |
rtac (mult_assoc RS trans) 3, |
|
1461 | 243 |
rtac (mult_commute RS subst_context) 6, |
244 |
REPEAT (ares_tac [mult_type] 1)]); |
|
437 | 245 |
|
246 |
val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; |
|
247 |
||
0 | 248 |
|
249 |
(*** Difference ***) |
|
250 |
||
760 | 251 |
qed_goal "diff_self_eq_0" Arith.thy |
0 | 252 |
"m:nat ==> m #- m = 0" |
253 |
(fn prems=> |
|
254 |
[ (nat_ind_tac "m" prems 1), |
|
4091 | 255 |
(ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); |
0 | 256 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
257 |
(*Addition is the inverse of subtraction*) |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
258 |
goal Arith.thy "!!m n. [| n le m; m:nat |] ==> n #+ (m#-n) = m"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
259 |
by (forward_tac [lt_nat_in_nat] 1); |
127 | 260 |
by (etac nat_succI 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
261 |
by (etac rev_mp 1); |
0 | 262 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
2469 | 263 |
by (ALLGOALS Asm_simp_tac); |
760 | 264 |
qed "add_diff_inverse"; |
0 | 265 |
|
1609 | 266 |
(*Proof is IDENTICAL to that above*) |
267 |
goal Arith.thy "!!m n. [| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"; |
|
268 |
by (forward_tac [lt_nat_in_nat] 1); |
|
269 |
by (etac nat_succI 1); |
|
270 |
by (etac rev_mp 1); |
|
271 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
2469 | 272 |
by (ALLGOALS Asm_simp_tac); |
1609 | 273 |
qed "diff_succ"; |
274 |
||
1708 | 275 |
(** Subtraction is the inverse of addition. **) |
276 |
||
0 | 277 |
val [mnat,nnat] = goal Arith.thy |
437 | 278 |
"[| m:nat; n:nat |] ==> (n#+m) #- n = m"; |
0 | 279 |
by (rtac (nnat RS nat_induct) 1); |
4091 | 280 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [mnat]))); |
760 | 281 |
qed "diff_add_inverse"; |
0 | 282 |
|
437 | 283 |
goal Arith.thy |
284 |
"!!m n. [| m:nat; n:nat |] ==> (m#+n) #- n = m"; |
|
285 |
by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1); |
|
286 |
by (REPEAT (ares_tac [diff_add_inverse] 1)); |
|
760 | 287 |
qed "diff_add_inverse2"; |
437 | 288 |
|
1708 | 289 |
goal Arith.thy |
290 |
"!!k. [| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n"; |
|
291 |
by (nat_ind_tac "k" [] 1); |
|
2469 | 292 |
by (ALLGOALS Asm_simp_tac); |
1708 | 293 |
qed "diff_cancel"; |
294 |
||
295 |
goal Arith.thy |
|
296 |
"!!n. [| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n"; |
|
297 |
val add_commute_k = read_instantiate [("n","k")] add_commute; |
|
4091 | 298 |
by (asm_simp_tac (simpset() addsimps [add_commute_k, diff_cancel]) 1); |
1708 | 299 |
qed "diff_cancel2"; |
300 |
||
0 | 301 |
val [mnat,nnat] = goal Arith.thy |
302 |
"[| m:nat; n:nat |] ==> n #- (n#+m) = 0"; |
|
303 |
by (rtac (nnat RS nat_induct) 1); |
|
4091 | 304 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [mnat]))); |
760 | 305 |
qed "diff_add_0"; |
0 | 306 |
|
1708 | 307 |
(** Difference distributes over multiplication **) |
308 |
||
309 |
goal Arith.thy |
|
310 |
"!!m n. [| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)"; |
|
311 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
4091 | 312 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel]))); |
1708 | 313 |
qed "diff_mult_distrib" ; |
314 |
||
315 |
goal Arith.thy |
|
316 |
"!!m. [| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)"; |
|
317 |
val mult_commute_k = read_instantiate [("m","k")] mult_commute; |
|
4091 | 318 |
by (asm_simp_tac (simpset() addsimps |
1793 | 319 |
[mult_commute_k, diff_mult_distrib]) 1); |
1708 | 320 |
qed "diff_mult_distrib2" ; |
321 |
||
0 | 322 |
(*** Remainder ***) |
323 |
||
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
324 |
goal Arith.thy "!!m n. [| 0<n; n le m; m:nat |] ==> m #- n < m"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
325 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
0 | 326 |
by (etac rev_mp 1); |
327 |
by (etac rev_mp 1); |
|
328 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
4091 | 329 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_le_self,diff_succ_succ]))); |
760 | 330 |
qed "div_termination"; |
0 | 331 |
|
1461 | 332 |
val div_rls = (*for mod and div*) |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
333 |
nat_typechecks @ |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
334 |
[Ord_transrec_type, apply_type, div_termination RS ltD, if_type, |
435 | 335 |
nat_into_Ord, not_lt_iff_le RS iffD1]; |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
336 |
|
4091 | 337 |
val div_ss = (simpset()) addsimps [nat_into_Ord, div_termination RS ltD, |
2493 | 338 |
not_lt_iff_le RS iffD2]; |
0 | 339 |
|
340 |
(*Type checking depends upon termination!*) |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
341 |
goalw Arith.thy [mod_def] "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n : nat"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
342 |
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); |
760 | 343 |
qed "mod_type"; |
0 | 344 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
345 |
goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m mod n = m"; |
0 | 346 |
by (rtac (mod_def RS def_transrec RS trans) 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
347 |
by (asm_simp_tac div_ss 1); |
760 | 348 |
qed "mod_less"; |
0 | 349 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
350 |
goal Arith.thy "!!m n. [| 0<n; n le m; m:nat |] ==> m mod n = (m#-n) mod n"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
351 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
0 | 352 |
by (rtac (mod_def RS def_transrec RS trans) 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
353 |
by (asm_simp_tac div_ss 1); |
760 | 354 |
qed "mod_geq"; |
0 | 355 |
|
2469 | 356 |
Addsimps [mod_type, mod_less, mod_geq]; |
357 |
||
0 | 358 |
(*** Quotient ***) |
359 |
||
360 |
(*Type checking depends upon termination!*) |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
361 |
goalw Arith.thy [div_def] |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
362 |
"!!m n. [| 0<n; m:nat; n:nat |] ==> m div n : nat"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
363 |
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); |
760 | 364 |
qed "div_type"; |
0 | 365 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
366 |
goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m div n = 0"; |
0 | 367 |
by (rtac (div_def RS def_transrec RS trans) 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
368 |
by (asm_simp_tac div_ss 1); |
760 | 369 |
qed "div_less"; |
0 | 370 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
371 |
goal Arith.thy |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
372 |
"!!m n. [| 0<n; n le m; m:nat |] ==> m div n = succ((m#-n) div n)"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
373 |
by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); |
0 | 374 |
by (rtac (div_def RS def_transrec RS trans) 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
375 |
by (asm_simp_tac div_ss 1); |
760 | 376 |
qed "div_geq"; |
0 | 377 |
|
2469 | 378 |
Addsimps [div_type, div_less, div_geq]; |
379 |
||
1609 | 380 |
(*A key result*) |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
381 |
goal Arith.thy |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
382 |
"!!m n. [| 0<n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
383 |
by (etac complete_induct 1); |
437 | 384 |
by (excluded_middle_tac "x<n" 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
385 |
(*case x<n*) |
2469 | 386 |
by (Asm_simp_tac 2); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
387 |
(*case n le x*) |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
388 |
by (asm_full_simp_tac |
4091 | 389 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, add_assoc, |
1461 | 390 |
div_termination RS ltD, add_diff_inverse]) 1); |
760 | 391 |
qed "mod_div_equality"; |
0 | 392 |
|
1609 | 393 |
(*** Further facts about mod (mainly for mutilated checkerboard ***) |
394 |
||
395 |
goal Arith.thy |
|
396 |
"!!m n. [| 0<n; m:nat; n:nat |] ==> \ |
|
397 |
\ succ(m) mod n = if(succ(m mod n) = n, 0, succ(m mod n))"; |
|
398 |
by (etac complete_induct 1); |
|
399 |
by (excluded_middle_tac "succ(x)<n" 1); |
|
1623 | 400 |
(* case succ(x) < n *) |
4091 | 401 |
by (asm_simp_tac (simpset() addsimps [mod_less, nat_le_refl RS lt_trans, |
1623 | 402 |
succ_neq_self]) 2); |
4091 | 403 |
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq]) 2); |
1623 | 404 |
(* case n le succ(x) *) |
1609 | 405 |
by (asm_full_simp_tac |
4091 | 406 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, mod_geq]) 1); |
1623 | 407 |
by (etac leE 1); |
4091 | 408 |
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD, diff_succ, |
1623 | 409 |
mod_geq]) 1); |
4091 | 410 |
by (asm_simp_tac (simpset() addsimps [mod_less, diff_self_eq_0]) 1); |
1609 | 411 |
qed "mod_succ"; |
412 |
||
413 |
goal Arith.thy "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n < n"; |
|
414 |
by (etac complete_induct 1); |
|
415 |
by (excluded_middle_tac "x<n" 1); |
|
416 |
(*case x<n*) |
|
4091 | 417 |
by (asm_simp_tac (simpset() addsimps [mod_less]) 2); |
1609 | 418 |
(*case n le x*) |
419 |
by (asm_full_simp_tac |
|
4091 | 420 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
1609 | 421 |
mod_geq, div_termination RS ltD]) 1); |
422 |
qed "mod_less_divisor"; |
|
423 |
||
424 |
||
425 |
goal Arith.thy |
|
426 |
"!!k b. [| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = if(b=1,0,1)"; |
|
427 |
by (subgoal_tac "k mod 2: 2" 1); |
|
4091 | 428 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
1623 | 429 |
by (dtac ltD 1); |
4091 | 430 |
by (asm_simp_tac (simpset() setloop split_tac [expand_if]) 1); |
3016 | 431 |
by (Blast_tac 1); |
1609 | 432 |
qed "mod2_cases"; |
433 |
||
434 |
goal Arith.thy "!!m. m:nat ==> succ(succ(m)) mod 2 = m mod 2"; |
|
435 |
by (subgoal_tac "m mod 2: 2" 1); |
|
4091 | 436 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
437 |
by (asm_simp_tac (simpset() addsimps [mod_succ] setloop Step_tac) 1); |
|
1609 | 438 |
qed "mod2_succ_succ"; |
439 |
||
440 |
goal Arith.thy "!!m. m:nat ==> (m#+m) mod 2 = 0"; |
|
1623 | 441 |
by (etac nat_induct 1); |
4091 | 442 |
by (simp_tac (simpset() addsimps [mod_less]) 1); |
443 |
by (asm_simp_tac (simpset() addsimps [mod2_succ_succ, add_succ_right]) 1); |
|
1609 | 444 |
qed "mod2_add_self"; |
445 |
||
0 | 446 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
447 |
(**** Additional theorems about "le" ****) |
0 | 448 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
449 |
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le m #+ n"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
450 |
by (etac nat_induct 1); |
2469 | 451 |
by (ALLGOALS Asm_simp_tac); |
760 | 452 |
qed "add_le_self"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
453 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
454 |
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le n #+ m"; |
2033 | 455 |
by (stac add_commute 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
456 |
by (REPEAT (ares_tac [add_le_self] 1)); |
760 | 457 |
qed "add_le_self2"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
458 |
|
1708 | 459 |
(*** Monotonicity of Addition ***) |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
460 |
|
1708 | 461 |
(*strict, in 1st argument; proof is by rule induction on 'less than'*) |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
462 |
goal Arith.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> i#+k < j#+k"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
463 |
by (forward_tac [lt_nat_in_nat] 1); |
127 | 464 |
by (assume_tac 1); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
465 |
by (etac succ_lt_induct 1); |
4091 | 466 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [leI]))); |
760 | 467 |
qed "add_lt_mono1"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
468 |
|
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
469 |
(*strict, in both arguments*) |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
470 |
goal Arith.thy "!!i j k l. [| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
471 |
by (rtac (add_lt_mono1 RS lt_trans) 1); |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
472 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); |
2033 | 473 |
by (EVERY [stac add_commute 1, |
474 |
stac add_commute 3, |
|
1461 | 475 |
rtac add_lt_mono1 5]); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
476 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); |
760 | 477 |
qed "add_lt_mono"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
478 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
479 |
(*A [clumsy] way of lifting < monotonicity to le monotonicity *) |
435 | 480 |
val lt_mono::ford::prems = goal Ordinal.thy |
1461 | 481 |
"[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \ |
482 |
\ !!i. i:k ==> Ord(f(i)); \ |
|
483 |
\ i le j; j:k \ |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
484 |
\ |] ==> f(i) le f(j)"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
485 |
by (cut_facts_tac prems 1); |
3016 | 486 |
by (blast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1); |
760 | 487 |
qed "Ord_lt_mono_imp_le_mono"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
488 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
489 |
(*le monotonicity, 1st argument*) |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
490 |
goal Arith.thy |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
491 |
"!!i j k. [| i le j; j:nat; k:nat |] ==> i#+k le j#+k"; |
3840 | 492 |
by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1); |
435 | 493 |
by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1)); |
760 | 494 |
qed "add_le_mono1"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
495 |
|
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
496 |
(* le monotonicity, BOTH arguments*) |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
497 |
goal Arith.thy |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
498 |
"!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l"; |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
499 |
by (rtac (add_le_mono1 RS le_trans) 1); |
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
500 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
2033 | 501 |
by (EVERY [stac add_commute 1, |
502 |
stac add_commute 3, |
|
1461 | 503 |
rtac add_le_mono1 5]); |
25
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents:
14
diff
changeset
|
504 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
760 | 505 |
qed "add_le_mono"; |
1609 | 506 |
|
1708 | 507 |
(*** Monotonicity of Multiplication ***) |
508 |
||
509 |
goal Arith.thy "!!i j k. [| i le j; j:nat; k:nat |] ==> i#*k le j#*k"; |
|
510 |
by (forward_tac [lt_nat_in_nat] 1); |
|
511 |
by (nat_ind_tac "k" [] 2); |
|
4091 | 512 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono]))); |
1708 | 513 |
qed "mult_le_mono1"; |
514 |
||
515 |
(* le monotonicity, BOTH arguments*) |
|
516 |
goal Arith.thy |
|
517 |
"!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l"; |
|
518 |
by (rtac (mult_le_mono1 RS le_trans) 1); |
|
519 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
|
2033 | 520 |
by (EVERY [stac mult_commute 1, |
521 |
stac mult_commute 3, |
|
1708 | 522 |
rtac mult_le_mono1 5]); |
523 |
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); |
|
524 |
qed "mult_le_mono"; |
|
525 |
||
526 |
(*strict, in 1st argument; proof is by induction on k>0*) |
|
527 |
goal Arith.thy "!!i j k. [| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j"; |
|
1793 | 528 |
by (etac zero_lt_natE 1); |
1708 | 529 |
by (forward_tac [lt_nat_in_nat] 2); |
2469 | 530 |
by (ALLGOALS Asm_simp_tac); |
1708 | 531 |
by (nat_ind_tac "x" [] 1); |
4091 | 532 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono]))); |
1708 | 533 |
qed "mult_lt_mono2"; |
534 |
||
535 |
goal Arith.thy "!!k. [| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n"; |
|
4091 | 536 |
by (best_tac (claset() addEs [natE] addss (simpset())) 1); |
1708 | 537 |
qed "zero_lt_mult_iff"; |
538 |
||
1793 | 539 |
goal Arith.thy "!!k. [| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1"; |
4091 | 540 |
by (best_tac (claset() addEs [natE] addss (simpset())) 1); |
1793 | 541 |
qed "mult_eq_1_iff"; |
542 |
||
1708 | 543 |
(*Cancellation law for division*) |
544 |
goal Arith.thy |
|
545 |
"!!k. [| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n"; |
|
546 |
by (eres_inst_tac [("i","m")] complete_induct 1); |
|
547 |
by (excluded_middle_tac "x<n" 1); |
|
4091 | 548 |
by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, |
1793 | 549 |
mult_lt_mono2]) 2); |
1708 | 550 |
by (asm_full_simp_tac |
4091 | 551 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
1708 | 552 |
zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq, |
553 |
diff_mult_distrib2 RS sym, |
|
1793 | 554 |
div_termination RS ltD]) 1); |
1708 | 555 |
qed "div_cancel"; |
556 |
||
557 |
goal Arith.thy |
|
558 |
"!!k. [| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> \ |
|
559 |
\ (k#*m) mod (k#*n) = k #* (m mod n)"; |
|
560 |
by (eres_inst_tac [("i","m")] complete_induct 1); |
|
561 |
by (excluded_middle_tac "x<n" 1); |
|
4091 | 562 |
by (asm_simp_tac (simpset() addsimps [mod_less, zero_lt_mult_iff, |
1793 | 563 |
mult_lt_mono2]) 2); |
1708 | 564 |
by (asm_full_simp_tac |
4091 | 565 |
(simpset() addsimps [not_lt_iff_le, nat_into_Ord, |
1708 | 566 |
zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq, |
567 |
diff_mult_distrib2 RS sym, |
|
1793 | 568 |
div_termination RS ltD]) 1); |
1708 | 569 |
qed "mult_mod_distrib"; |
570 |
||
1793 | 571 |
(** Lemma for gcd **) |
1708 | 572 |
|
1793 | 573 |
val mono_lemma = (nat_into_Ord RS Ord_0_lt) RSN (2,mult_lt_mono2); |
574 |
||
575 |
goal Arith.thy "!!m n. [| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0"; |
|
576 |
by (rtac disjCI 1); |
|
577 |
by (dtac sym 1); |
|
578 |
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I])); |
|
4091 | 579 |
by (fast_tac (claset() addss (simpset())) 1); |
2469 | 580 |
by (fast_tac (le_cs addDs [mono_lemma] |
4091 | 581 |
addss (simpset() addsimps [mult_1_right])) 1); |
1793 | 582 |
qed "mult_eq_self_implies_10"; |
1708 | 583 |
|
2469 | 584 |
|
585 |
(*Thanks to Sten Agerholm*) |
|
586 |
goal Arith.thy (* add_le_elim1 *) |
|
587 |
"!!m n k. [|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k"; |
|
2493 | 588 |
by (etac rev_mp 1); |
3016 | 589 |
by (eres_inst_tac [("n","n")] nat_induct 1); |
2469 | 590 |
by (Asm_simp_tac 1); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3207
diff
changeset
|
591 |
by Safe_tac; |
4091 | 592 |
by (asm_full_simp_tac (simpset() addsimps [not_le_iff_lt,nat_into_Ord]) 1); |
2469 | 593 |
by (etac lt_asym 1); |
594 |
by (assume_tac 1); |
|
595 |
by (Asm_full_simp_tac 1); |
|
4091 | 596 |
by (asm_full_simp_tac (simpset() addsimps [le_iff, nat_into_Ord]) 1); |
3016 | 597 |
by (Blast_tac 1); |
2469 | 598 |
qed "add_le_elim1"; |
599 |