author | wenzelm |
Mon, 21 Jan 2002 15:29:06 +0100 | |
changeset 12828 | 57fb9d1ee34a |
parent 12089 | 34e7693271a9 |
permissions | -rw-r--r-- |
9548 | 1 |
(* Title: ZF/ArithSimp.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 2000 University of Cambridge |
|
5 |
||
6 |
Arithmetic with simplification |
|
7 |
*) |
|
8 |
||
9 |
||
10 |
Addsimprocs ArithData.nat_cancel; |
|
11 |
||
12 |
||
13 |
(*** Difference ***) |
|
14 |
||
15 |
Goal "m #- m = 0"; |
|
16 |
by (subgoal_tac "natify(m) #- natify(m) = 0" 1); |
|
17 |
by (rtac (natify_in_nat RS nat_induct) 2); |
|
18 |
by Auto_tac; |
|
19 |
qed "diff_self_eq_0"; |
|
20 |
||
21 |
(**Addition is the inverse of subtraction**) |
|
22 |
||
23 |
(*We need m:nat even if we replace the RHS by natify(m), for consider e.g. |
|
24 |
n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 ~= 0 = natify(m).*) |
|
25 |
Goal "[| n le m; m:nat |] ==> n #+ (m#-n) = m"; |
|
26 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
27 |
by (etac rev_mp 1); |
|
28 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
29 |
by Auto_tac; |
|
30 |
qed "add_diff_inverse"; |
|
31 |
||
32 |
Goal "[| n le m; m:nat |] ==> (m#-n) #+ n = m"; |
|
33 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
34 |
by (asm_simp_tac (simpset() addsimps [add_commute, add_diff_inverse]) 1); |
|
35 |
qed "add_diff_inverse2"; |
|
36 |
||
37 |
(*Proof is IDENTICAL to that of add_diff_inverse*) |
|
38 |
Goal "[| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"; |
|
39 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
40 |
by (etac rev_mp 1); |
|
41 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
42 |
by (ALLGOALS Asm_simp_tac); |
|
43 |
qed "diff_succ"; |
|
44 |
||
45 |
Goal "[| m: nat; n: nat |] ==> 0 < (n #- m) <-> m<n"; |
|
46 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
47 |
by (ALLGOALS Asm_simp_tac); |
|
48 |
qed "zero_less_diff"; |
|
49 |
Addsimps [zero_less_diff]; |
|
50 |
||
51 |
||
52 |
(** Difference distributes over multiplication **) |
|
53 |
||
54 |
Goal "(m #- n) #* k = (m #* k) #- (n #* k)"; |
|
55 |
by (subgoal_tac "(natify(m) #- natify(n)) #* natify(k) = \ |
|
56 |
\ (natify(m) #* natify(k)) #- (natify(n) #* natify(k))" 1); |
|
57 |
by (res_inst_tac [("m","natify(m)"),("n","natify(n)")] diff_induct 2); |
|
58 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [diff_cancel]))); |
|
59 |
qed "diff_mult_distrib" ; |
|
60 |
||
61 |
Goal "k #* (m #- n) = (k #* m) #- (k #* n)"; |
|
62 |
by (simp_tac |
|
63 |
(simpset() addsimps [inst "m" "k" mult_commute, diff_mult_distrib]) 1); |
|
64 |
qed "diff_mult_distrib2"; |
|
65 |
||
66 |
||
67 |
(*** Remainder ***) |
|
68 |
||
69 |
(*We need m:nat even with natify*) |
|
70 |
Goal "[| 0<n; n le m; m:nat |] ==> m #- n < m"; |
|
71 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
72 |
by (etac rev_mp 1); |
|
73 |
by (etac rev_mp 1); |
|
74 |
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
75 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_le_self]))); |
|
76 |
qed "div_termination"; |
|
77 |
||
9907 | 78 |
bind_thms ("div_rls", (*for mod and div*) |
9548 | 79 |
nat_typechecks @ |
80 |
[Ord_transrec_type, apply_funtype, div_termination RS ltD, |
|
9907 | 81 |
nat_into_Ord, not_lt_iff_le RS iffD1]); |
9548 | 82 |
|
83 |
val div_ss = simpset() addsimps [div_termination RS ltD, |
|
84 |
not_lt_iff_le RS iffD2]; |
|
85 |
||
86 |
Goalw [raw_mod_def] "[| m:nat; n:nat |] ==> raw_mod (m, n) : nat"; |
|
87 |
by (rtac Ord_transrec_type 1); |
|
88 |
by (auto_tac(claset(), simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff])); |
|
89 |
by (REPEAT (ares_tac div_rls 1)); |
|
90 |
qed "raw_mod_type"; |
|
91 |
||
92 |
Goalw [mod_def] "m mod n : nat"; |
|
93 |
by (simp_tac (simpset() addsimps [mod_def, raw_mod_type]) 1); |
|
94 |
qed "mod_type"; |
|
95 |
AddTCs [mod_type]; |
|
96 |
AddIffs [mod_type]; |
|
97 |
||
98 |
||
99 |
(** Aribtrary definitions for division by zero. Useful to simplify |
|
100 |
certain equations **) |
|
101 |
||
102 |
Goalw [div_def] "a div 0 = 0"; |
|
103 |
by (rtac (raw_div_def RS def_transrec RS trans) 1); |
|
104 |
by (Asm_simp_tac 1); |
|
105 |
qed "DIVISION_BY_ZERO_DIV"; (*NOT for adding to default simpset*) |
|
106 |
||
107 |
Goalw [mod_def] "a mod 0 = natify(a)"; |
|
108 |
by (rtac (raw_mod_def RS def_transrec RS trans) 1); |
|
109 |
by (Asm_simp_tac 1); |
|
110 |
qed "DIVISION_BY_ZERO_MOD"; (*NOT for adding to default simpset*) |
|
111 |
||
112 |
fun div_undefined_case_tac s i = |
|
113 |
case_tac s i THEN |
|
114 |
asm_full_simp_tac |
|
115 |
(simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff]) (i+1) THEN |
|
11386 | 116 |
asm_full_simp_tac (simpset() addsimps [DIVISION_BY_ZERO_DIV, |
117 |
DIVISION_BY_ZERO_MOD]) i; |
|
9548 | 118 |
|
119 |
Goal "m<n ==> raw_mod (m,n) = m"; |
|
120 |
by (rtac (raw_mod_def RS def_transrec RS trans) 1); |
|
121 |
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD]) 1); |
|
122 |
qed "raw_mod_less"; |
|
123 |
||
124 |
Goal "[| m<n; n : nat |] ==> m mod n = m"; |
|
125 |
by (ftac lt_nat_in_nat 1 THEN assume_tac 1); |
|
126 |
by (asm_simp_tac (simpset() addsimps [mod_def, raw_mod_less]) 1); |
|
127 |
qed "mod_less"; |
|
128 |
||
129 |
Goal "[| 0<n; n le m; m:nat |] ==> raw_mod (m, n) = raw_mod (m#-n, n)"; |
|
130 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
131 |
by (rtac (raw_mod_def RS def_transrec RS trans) 1); |
|
132 |
by (asm_simp_tac div_ss 1); |
|
133 |
by (Blast_tac 1); |
|
134 |
qed "raw_mod_geq"; |
|
135 |
||
136 |
Goal "[| n le m; m:nat |] ==> m mod n = (m#-n) mod n"; |
|
137 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
138 |
by (div_undefined_case_tac "n=0" 1); |
|
139 |
by (asm_simp_tac (simpset() addsimps [mod_def, raw_mod_geq]) 1); |
|
140 |
qed "mod_geq"; |
|
141 |
||
142 |
Addsimps [mod_less]; |
|
143 |
||
144 |
||
145 |
(*** Division ***) |
|
146 |
||
147 |
Goalw [raw_div_def] "[| m:nat; n:nat |] ==> raw_div (m, n) : nat"; |
|
148 |
by (rtac Ord_transrec_type 1); |
|
149 |
by (auto_tac(claset(), simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff])); |
|
150 |
by (REPEAT (ares_tac div_rls 1)); |
|
151 |
qed "raw_div_type"; |
|
152 |
||
153 |
Goalw [div_def] "m div n : nat"; |
|
154 |
by (simp_tac (simpset() addsimps [div_def, raw_div_type]) 1); |
|
155 |
qed "div_type"; |
|
156 |
AddTCs [div_type]; |
|
157 |
AddIffs [div_type]; |
|
158 |
||
159 |
Goal "m<n ==> raw_div (m,n) = 0"; |
|
160 |
by (rtac (raw_div_def RS def_transrec RS trans) 1); |
|
161 |
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD]) 1); |
|
162 |
qed "raw_div_less"; |
|
163 |
||
164 |
Goal "[| m<n; n : nat |] ==> m div n = 0"; |
|
165 |
by (ftac lt_nat_in_nat 1 THEN assume_tac 1); |
|
166 |
by (asm_simp_tac (simpset() addsimps [div_def, raw_div_less]) 1); |
|
167 |
qed "div_less"; |
|
168 |
||
169 |
Goal "[| 0<n; n le m; m:nat |] ==> raw_div(m,n) = succ(raw_div(m#-n, n))"; |
|
170 |
by (subgoal_tac "n ~= 0" 1); |
|
171 |
by (Blast_tac 2); |
|
172 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
173 |
by (rtac (raw_div_def RS def_transrec RS trans) 1); |
|
174 |
by (asm_simp_tac div_ss 1); |
|
175 |
qed "raw_div_geq"; |
|
176 |
||
177 |
Goal "[| 0<n; n le m; m:nat |] ==> m div n = succ ((m#-n) div n)"; |
|
178 |
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1); |
|
179 |
by (asm_simp_tac (simpset() addsimps [div_def, raw_div_geq]) 1); |
|
180 |
qed "div_geq"; |
|
181 |
||
182 |
Addsimps [div_less, div_geq]; |
|
183 |
||
184 |
||
185 |
(*A key result*) |
|
186 |
Goal "[| m: nat; n: nat |] ==> (m div n)#*n #+ m mod n = m"; |
|
187 |
by (div_undefined_case_tac "n=0" 1); |
|
188 |
by (etac complete_induct 1); |
|
189 |
by (case_tac "x<n" 1); |
|
190 |
(*case n le x*) |
|
191 |
by (asm_full_simp_tac |
|
192 |
(simpset() addsimps [not_lt_iff_le, add_assoc, mod_geq, |
|
193 |
div_termination RS ltD, add_diff_inverse]) 2); |
|
194 |
(*case x<n*) |
|
195 |
by (Asm_simp_tac 1); |
|
196 |
val lemma = result(); |
|
197 |
||
198 |
Goal "(m div n)#*n #+ m mod n = natify(m)"; |
|
199 |
by (subgoal_tac |
|
200 |
"(natify(m) div natify(n))#*natify(n) #+ natify(m) mod natify(n) = \ |
|
201 |
\ natify(m)" 1); |
|
202 |
by (stac lemma 2); |
|
203 |
by Auto_tac; |
|
204 |
qed "mod_div_equality_natify"; |
|
205 |
||
206 |
Goal "m: nat ==> (m div n)#*n #+ m mod n = m"; |
|
207 |
by (asm_simp_tac (simpset() addsimps [mod_div_equality_natify]) 1); |
|
208 |
qed "mod_div_equality"; |
|
209 |
||
210 |
||
211 |
(*** Further facts about mod (mainly for mutilated chess board) ***) |
|
212 |
||
213 |
Goal "[| 0<n; m:nat; n:nat |] \ |
|
214 |
\ ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"; |
|
215 |
by (etac complete_induct 1); |
|
216 |
by (excluded_middle_tac "succ(x)<n" 1); |
|
217 |
(* case succ(x) < n *) |
|
218 |
by (asm_simp_tac (simpset() addsimps [nat_le_refl RS lt_trans, |
|
219 |
succ_neq_self]) 2); |
|
220 |
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq]) 2); |
|
221 |
(* case n le succ(x) *) |
|
222 |
by (asm_full_simp_tac (simpset() addsimps [mod_geq, not_lt_iff_le]) 1); |
|
223 |
by (etac leE 1); |
|
224 |
(*equality case*) |
|
225 |
by (asm_full_simp_tac (simpset() addsimps [diff_self_eq_0]) 2); |
|
226 |
by (asm_simp_tac (simpset() addsimps [mod_geq, div_termination RS ltD, |
|
227 |
diff_succ]) 1); |
|
228 |
val lemma = result(); |
|
229 |
||
230 |
Goal "n:nat ==> \ |
|
231 |
\ succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"; |
|
232 |
by (case_tac "n=0" 1); |
|
233 |
by (asm_simp_tac (simpset() addsimps [natify_succ, DIVISION_BY_ZERO_MOD]) 1); |
|
234 |
by (subgoal_tac |
|
235 |
"natify(succ(m)) mod n = \ |
|
236 |
\ (if succ(natify(m) mod n) = n then 0 else succ(natify(m) mod n))" 1); |
|
237 |
by (stac natify_succ 2); |
|
238 |
by (rtac lemma 2); |
|
239 |
by (auto_tac(claset(), |
|
240 |
simpset() delsimps [natify_succ] |
|
241 |
addsimps [nat_into_Ord RS Ord_0_lt_iff])); |
|
242 |
qed "mod_succ"; |
|
243 |
||
244 |
Goal "[| 0<n; n:nat |] ==> m mod n < n"; |
|
245 |
by (subgoal_tac "natify(m) mod n < n" 1); |
|
246 |
by (res_inst_tac [("i","natify(m)")] complete_induct 2); |
|
247 |
by (case_tac "x<n" 3); |
|
248 |
(*case x<n*) |
|
249 |
by (Asm_simp_tac 3); |
|
250 |
(*case n le x*) |
|
251 |
by (asm_full_simp_tac |
|
252 |
(simpset() addsimps [mod_geq, not_lt_iff_le, div_termination RS ltD]) 3); |
|
253 |
by Auto_tac; |
|
254 |
qed "mod_less_divisor"; |
|
255 |
||
11386 | 256 |
Goal "m mod 1 = 0"; |
257 |
by (cut_inst_tac [("n","1")] mod_less_divisor 1); |
|
258 |
by Auto_tac; |
|
259 |
qed "mod_1_eq"; |
|
260 |
Addsimps [mod_1_eq]; |
|
261 |
||
9548 | 262 |
Goal "b<2 ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"; |
263 |
by (subgoal_tac "k mod 2: 2" 1); |
|
264 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
|
265 |
by (dtac ltD 1); |
|
266 |
by Auto_tac; |
|
267 |
qed "mod2_cases"; |
|
268 |
||
269 |
Goal "succ(succ(m)) mod 2 = m mod 2"; |
|
270 |
by (subgoal_tac "m mod 2: 2" 1); |
|
271 |
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); |
|
272 |
by (auto_tac (claset(), simpset() addsimps [mod_succ])); |
|
273 |
qed "mod2_succ_succ"; |
|
274 |
||
275 |
Addsimps [mod2_succ_succ]; |
|
276 |
||
277 |
Goal "(m#+m#+n) mod 2 = n mod 2"; |
|
278 |
by (subgoal_tac "(natify(m)#+natify(m)#+n) mod 2 = n mod 2" 1); |
|
279 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
|
280 |
by Auto_tac; |
|
281 |
qed "mod2_add_more"; |
|
282 |
||
283 |
Goal "(m#+m) mod 2 = 0"; |
|
284 |
by (cut_inst_tac [("n","0")] mod2_add_more 1); |
|
285 |
by Auto_tac; |
|
286 |
qed "mod2_add_self"; |
|
287 |
||
288 |
Addsimps [mod2_add_more, mod2_add_self]; |
|
289 |
||
290 |
||
291 |
(**** Additional theorems about "le" ****) |
|
292 |
||
293 |
Goal "m:nat ==> m le (m #+ n)"; |
|
9873 | 294 |
by (Asm_simp_tac 1); |
9548 | 295 |
qed "add_le_self"; |
296 |
||
297 |
Goal "m:nat ==> m le (n #+ m)"; |
|
9873 | 298 |
by (Asm_simp_tac 1); |
9548 | 299 |
qed "add_le_self2"; |
300 |
||
301 |
(*** Monotonicity of Multiplication ***) |
|
302 |
||
303 |
Goal "[| i le j; j:nat |] ==> (i#*k) le (j#*k)"; |
|
304 |
by (subgoal_tac "natify(i)#*natify(k) le j#*natify(k)" 1); |
|
305 |
by (ftac lt_nat_in_nat 2); |
|
306 |
by (res_inst_tac [("n","natify(k)")] nat_induct 3); |
|
307 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_le_mono]))); |
|
308 |
qed "mult_le_mono1"; |
|
309 |
||
310 |
(* le monotonicity, BOTH arguments*) |
|
311 |
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l"; |
|
312 |
by (rtac (mult_le_mono1 RS le_trans) 1); |
|
313 |
by (REPEAT (assume_tac 1)); |
|
314 |
by (EVERY [stac mult_commute 1, |
|
315 |
stac mult_commute 1, |
|
316 |
rtac mult_le_mono1 1]); |
|
317 |
by (REPEAT (assume_tac 1)); |
|
318 |
qed "mult_le_mono"; |
|
319 |
||
320 |
(*strict, in 1st argument; proof is by induction on k>0. |
|
321 |
I can't see how to relax the typing conditions.*) |
|
322 |
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j"; |
|
323 |
by (etac zero_lt_natE 1); |
|
324 |
by (ftac lt_nat_in_nat 2); |
|
325 |
by (ALLGOALS Asm_simp_tac); |
|
326 |
by (induct_tac "x" 1); |
|
327 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono]))); |
|
328 |
qed "mult_lt_mono2"; |
|
329 |
||
330 |
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> i#*k < j#*k"; |
|
331 |
by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, |
|
332 |
inst "n" "k" mult_commute]) 1); |
|
333 |
qed "mult_lt_mono1"; |
|
334 |
||
335 |
Goal "m#+n = 0 <-> natify(m)=0 & natify(n)=0"; |
|
336 |
by (subgoal_tac "natify(m) #+ natify(n) = 0 <-> natify(m)=0 & natify(n)=0" 1); |
|
337 |
by (res_inst_tac [("n","natify(m)")] natE 2); |
|
338 |
by (res_inst_tac [("n","natify(n)")] natE 4); |
|
339 |
by Auto_tac; |
|
340 |
qed "add_eq_0_iff"; |
|
341 |
AddIffs [add_eq_0_iff]; |
|
342 |
||
343 |
Goal "0 < m#*n <-> 0 < natify(m) & 0 < natify(n)"; |
|
344 |
by (subgoal_tac "0 < natify(m)#*natify(n) <-> \ |
|
345 |
\ 0 < natify(m) & 0 < natify(n)" 1); |
|
346 |
by (res_inst_tac [("n","natify(m)")] natE 2); |
|
347 |
by (res_inst_tac [("n","natify(n)")] natE 4); |
|
348 |
by (res_inst_tac [("n","natify(n)")] natE 3); |
|
349 |
by Auto_tac; |
|
350 |
qed "zero_lt_mult_iff"; |
|
351 |
||
352 |
Goal "m#*n = 1 <-> natify(m)=1 & natify(n)=1"; |
|
353 |
by (subgoal_tac "natify(m) #* natify(n) = 1 <-> natify(m)=1 & natify(n)=1" 1); |
|
354 |
by (res_inst_tac [("n","natify(m)")] natE 2); |
|
355 |
by (res_inst_tac [("n","natify(n)")] natE 4); |
|
356 |
by Auto_tac; |
|
357 |
qed "mult_eq_1_iff"; |
|
358 |
AddIffs [zero_lt_mult_iff, mult_eq_1_iff]; |
|
359 |
||
11386 | 360 |
Goal "[|m: nat; n: nat|] ==> (m #* n = 0) <-> (m = 0 | n = 0)"; |
361 |
by Auto_tac; |
|
362 |
by (etac natE 1); |
|
363 |
by (etac natE 2); |
|
364 |
by Auto_tac; |
|
365 |
qed "mult_is_zero"; |
|
366 |
||
367 |
Goal "(m #* n = 0) <-> (natify(m) = 0 | natify(n) = 0)"; |
|
368 |
by (cut_inst_tac [("m","natify(m)"),("n","natify(n)")] mult_is_zero 1); |
|
369 |
by Auto_tac; |
|
370 |
qed "mult_is_zero_natify"; |
|
371 |
AddIffs [mult_is_zero_natify]; |
|
372 |
||
9648
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
373 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
374 |
(** Cancellation laws for common factors in comparisons **) |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
375 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
376 |
Goal "[| k: nat; m: nat; n: nat |] ==> (m#*k < n#*k) <-> (0<k & m<n)"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
377 |
by (safe_tac (claset() addSIs [mult_lt_mono1])); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
378 |
by (etac natE 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
379 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
380 |
by (rtac (not_le_iff_lt RS iffD1) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
381 |
by (dtac (not_le_iff_lt RSN (2,rev_iffD2)) 3); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
382 |
by (blast_tac (claset() addIs [mult_le_mono1]) 5); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
383 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
384 |
val lemma = result(); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
385 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
386 |
Goal "(m#*k < n#*k) <-> (0 < natify(k) & natify(m) < natify(n))"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
387 |
by (rtac iff_trans 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
388 |
by (rtac lemma 2); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
389 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
390 |
qed "mult_less_cancel2"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
391 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
392 |
Goal "(k#*m < k#*n) <-> (0 < natify(k) & natify(m) < natify(n))"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
393 |
by (simp_tac (simpset() addsimps [mult_less_cancel2, |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
394 |
inst "m" "k" mult_commute]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
395 |
qed "mult_less_cancel1"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
396 |
Addsimps [mult_less_cancel1, mult_less_cancel2]; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
397 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
398 |
Goal "(m#*k le n#*k) <-> (0 < natify(k) --> natify(m) le natify(n))"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
399 |
by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
400 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
401 |
qed "mult_le_cancel2"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
402 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
403 |
Goal "(k#*m le k#*n) <-> (0 < natify(k) --> natify(m) le natify(n))"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
404 |
by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
405 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
406 |
qed "mult_le_cancel1"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
407 |
Addsimps [mult_le_cancel1, mult_le_cancel2]; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
408 |
|
11386 | 409 |
Goal "k : nat ==> k #* m le k \\<longleftrightarrow> (0 < k \\<longrightarrow> natify(m) le 1)"; |
410 |
by (cut_inst_tac [("k","k"),("m","m"),("n","1")] mult_le_cancel1 1); |
|
411 |
by Auto_tac; |
|
412 |
qed "mult_le_cancel_le1"; |
|
413 |
||
9648
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
414 |
Goal "[| Ord(m); Ord(n) |] ==> m=n <-> (m le n & n le m)"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
415 |
by (blast_tac (claset() addIs [le_anti_sym]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
416 |
qed "Ord_eq_iff_le"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
417 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
418 |
Goal "[| k: nat; m: nat; n: nat |] ==> (m#*k = n#*k) <-> (m=n | k=0)"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
419 |
by (asm_simp_tac (simpset() addsimps [inst "m" "m#*k" Ord_eq_iff_le, |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
420 |
inst "m" "m" Ord_eq_iff_le]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
421 |
by (auto_tac (claset(), simpset() addsimps [Ord_0_lt_iff])); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
422 |
val lemma = result(); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
423 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
424 |
Goal "(m#*k = n#*k) <-> (natify(m) = natify(n) | natify(k) = 0)"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
425 |
by (rtac iff_trans 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
426 |
by (rtac lemma 2); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
427 |
by Auto_tac; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
428 |
qed "mult_cancel2"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
429 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
430 |
Goal "(k#*m = k#*n) <-> (natify(m) = natify(n) | natify(k) = 0)"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
431 |
by (simp_tac (simpset() addsimps [mult_cancel2, inst "m" "k" mult_commute]) 1); |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
432 |
qed "mult_cancel1"; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
433 |
Addsimps [mult_cancel1, mult_cancel2]; |
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
434 |
|
35d761c7d934
better rules for cancellation of common factors across comparisons
paulson
parents:
9548
diff
changeset
|
435 |
|
11386 | 436 |
(** Cancellation law for division **) |
437 |
||
9548 | 438 |
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n"; |
439 |
by (eres_inst_tac [("i","m")] complete_induct 1); |
|
440 |
by (excluded_middle_tac "x<n" 1); |
|
441 |
by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, |
|
442 |
mult_lt_mono2]) 2); |
|
443 |
by (asm_full_simp_tac |
|
444 |
(simpset() addsimps [not_lt_iff_le, |
|
445 |
zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq, |
|
446 |
diff_mult_distrib2 RS sym, |
|
447 |
div_termination RS ltD]) 1); |
|
11386 | 448 |
qed "div_cancel_raw"; |
449 |
||
450 |
Goal "[| 0 < natify(n); 0 < natify(k) |] ==> (k#*m) div (k#*n) = m div n"; |
|
451 |
by (cut_inst_tac [("k","natify(k)"),("m","natify(m)"),("n","natify(n)")] |
|
452 |
div_cancel_raw 1); |
|
453 |
by Auto_tac; |
|
9548 | 454 |
qed "div_cancel"; |
455 |
||
11386 | 456 |
|
457 |
(** Distributive law for remainder (mod) **) |
|
458 |
||
459 |
Goal "[| k:nat; m:nat; n:nat |] ==> (k#*m) mod (k#*n) = k #* (m mod n)"; |
|
460 |
by (div_undefined_case_tac "k=0" 1); |
|
461 |
by (div_undefined_case_tac "n=0" 1); |
|
9548 | 462 |
by (eres_inst_tac [("i","m")] complete_induct 1); |
11386 | 463 |
by (case_tac "x<n" 1); |
464 |
by (asm_simp_tac |
|
465 |
(simpset() addsimps [mod_less, zero_lt_mult_iff, mult_lt_mono2]) 1); |
|
9548 | 466 |
by (asm_full_simp_tac |
467 |
(simpset() addsimps [not_lt_iff_le, |
|
468 |
zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq, |
|
469 |
diff_mult_distrib2 RS sym, |
|
470 |
div_termination RS ltD]) 1); |
|
11386 | 471 |
qed "mult_mod_distrib_raw"; |
472 |
||
473 |
Goal "k #* (m mod n) = (k#*m) mod (k#*n)"; |
|
474 |
by (cut_inst_tac [("k","natify(k)"),("m","natify(m)"),("n","natify(n)")] |
|
475 |
mult_mod_distrib_raw 1); |
|
476 |
by Auto_tac; |
|
477 |
qed "mod_mult_distrib2"; |
|
478 |
||
479 |
Goal "(m mod n) #* k = (m#*k) mod (n#*k)"; |
|
480 |
by (simp_tac (simpset() addsimps [mult_commute, mod_mult_distrib2]) 1); |
|
9548 | 481 |
qed "mult_mod_distrib"; |
482 |
||
11386 | 483 |
Goal "n \\<in> nat ==> (m #+ n) mod n = m mod n"; |
484 |
by (subgoal_tac "(n #+ m) mod n = (n #+ m #- n) mod n" 1); |
|
485 |
by (stac (mod_geq RS sym) 2); |
|
486 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_commute]))); |
|
487 |
qed "mod_add_self2_raw"; |
|
488 |
||
489 |
Goal "(m #+ n) mod n = m mod n"; |
|
490 |
by (cut_inst_tac [("n","natify(n)")] mod_add_self2_raw 1); |
|
491 |
by Auto_tac; |
|
492 |
qed "mod_add_self2"; |
|
493 |
||
494 |
Goal "(n#+m) mod n = m mod n"; |
|
495 |
by (asm_simp_tac (simpset() addsimps [add_commute, mod_add_self2]) 1); |
|
496 |
qed "mod_add_self1"; |
|
497 |
Addsimps [mod_add_self1, mod_add_self2]; |
|
498 |
||
499 |
||
500 |
Goal "k \\<in> nat ==> (m #+ k#*n) mod n = m mod n"; |
|
501 |
by (etac nat_induct 1); |
|
502 |
by (ALLGOALS |
|
503 |
(asm_simp_tac |
|
504 |
(simpset() addsimps [inst "n" "n" add_left_commute]))); |
|
505 |
qed "mod_mult_self1_raw"; |
|
506 |
||
507 |
Goal "(m #+ k#*n) mod n = m mod n"; |
|
508 |
by (cut_inst_tac [("k","natify(k)")] mod_mult_self1_raw 1); |
|
509 |
by Auto_tac; |
|
510 |
qed "mod_mult_self1"; |
|
511 |
||
512 |
Goal "(m #+ n#*k) mod n = m mod n"; |
|
513 |
by (simp_tac (simpset() addsimps [mult_commute, mod_mult_self1]) 1); |
|
514 |
qed "mod_mult_self2"; |
|
515 |
||
516 |
Addsimps [mod_mult_self1, mod_mult_self2]; |
|
517 |
||
9548 | 518 |
(*Lemma for gcd*) |
519 |
Goal "m = m#*n ==> natify(n)=1 | m=0"; |
|
520 |
by (subgoal_tac "m: nat" 1); |
|
521 |
by (etac ssubst 2); |
|
522 |
by (rtac disjCI 1); |
|
523 |
by (dtac sym 1); |
|
524 |
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I])); |
|
525 |
by (dtac (nat_into_Ord RS Ord_0_lt RSN (2,mult_lt_mono2)) 3); |
|
526 |
by Auto_tac; |
|
527 |
by (subgoal_tac "m #* n = 0" 1); |
|
528 |
by (stac (mult_natify2 RS sym) 2); |
|
529 |
by (auto_tac (claset(), simpset() delsimps [mult_natify2])); |
|
530 |
qed "mult_eq_self_implies_10"; |
|
531 |
||
532 |
Goal "[| m<n; n: nat |] ==> EX k: nat. n = succ(m#+k)"; |
|
533 |
by (ftac lt_nat_in_nat 1 THEN assume_tac 1); |
|
534 |
by (etac rev_mp 1); |
|
535 |
by (induct_tac "n" 1); |
|
536 |
by (ALLGOALS (simp_tac (simpset() addsimps [le_iff]))); |
|
537 |
by (blast_tac (claset() addSEs [leE] |
|
538 |
addSIs [add_0_right RS sym, add_succ_right RS sym]) 1); |
|
539 |
qed_spec_mp "less_imp_succ_add"; |
|
540 |
||
9883 | 541 |
Goal "[| m: nat; n: nat |] ==> (m<n) <-> (EX k: nat. n = succ(m#+k))"; |
542 |
by (auto_tac (claset() addIs [less_imp_succ_add], simpset())); |
|
543 |
qed "less_iff_succ_add"; |
|
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
544 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
545 |
(* on nat *) |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
546 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
547 |
Goal "m #- n = 0 <-> natify(m) le natify(n)"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
548 |
by (auto_tac (claset(), simpset() addsimps |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
549 |
[le_iff, diff_self_eq_0])); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
550 |
by (full_simp_tac (simpset() addsimps [less_iff_succ_add ]) 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
551 |
by (Clarify_tac 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
552 |
by (subgoal_tac "natify(m) #- natify(n) = 0" 3); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
553 |
by (etac subst 3); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
554 |
by (ALLGOALS(asm_full_simp_tac (simpset() |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
555 |
addsimps [diff_self_eq_0]))); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
556 |
by (subgoal_tac "natify(m) #- natify(n) = 0" 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
557 |
by (etac subst 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
558 |
by (etac ssubst 3); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
559 |
by (rtac (not_le_iff_lt RS iffD1) 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
560 |
by (auto_tac (claset(), simpset() addsimps [le_iff])); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
561 |
by (subgoal_tac "natify(m) #- natify(n) = 0" 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
562 |
by (Asm_simp_tac 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
563 |
by (thin_tac "m #- n = 0" 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
564 |
by (dtac ([natify_in_nat, natify_in_nat] MRS zero_less_diff RS iffD2) 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
565 |
by Auto_tac; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
566 |
qed "diff_is_0_iff"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
567 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
568 |
Goal "[| a:nat; b:nat |] ==> \ |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
569 |
\ (P(a #- b)) <-> ((a < b -->P(0)) & (ALL d:nat. a = b #+ d --> P(d)))"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
570 |
by (case_tac "a le b" 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
571 |
by (subgoal_tac "natify(a) le natify(b)" 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
572 |
by (dtac (diff_is_0_iff RS iffD2) 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
573 |
by Safe_tac; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
574 |
by (ALLGOALS(Asm_full_simp_tac)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
575 |
by (ALLGOALS(rotate_tac ~1)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
576 |
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [le_iff]))); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
577 |
by (Clarify_tac 2); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
578 |
by (ALLGOALS(dtac not_lt_imp_le)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
579 |
by (ALLGOALS(Asm_full_simp_tac)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
580 |
by (ALLGOALS(dres_inst_tac [("x", "a #- b")] bspec)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
581 |
by (ALLGOALS(Asm_simp_tac)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
582 |
by (ALLGOALS(dtac add_diff_inverse)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
583 |
by (ALLGOALS(rotate_tac ~1)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
584 |
by (ALLGOALS(Asm_full_simp_tac)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
585 |
qed "nat_diff_split"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
586 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
587 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
11386
diff
changeset
|
588 |