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