author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
parent 7549 | 1dcf2a7a2b5b |
child 8257 | fe9bf28e8a58 |
permissions | -rw-r--r-- |
6917 | 1 |
(* Title: HOL/IntDiv.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
6 |
The division operators div, mod and the divides relation "dvd" |
|
6992 | 7 |
|
8 |
Here is the division algorithm in ML: |
|
9 |
||
10 |
fun posDivAlg (a,b) = |
|
11 |
if a<b then (0,a) |
|
12 |
else let val (q,r) = posDivAlg(a, 2*b) |
|
13 |
in if 0<=r-b then (2*q+1, r-b) else (2*q, r) |
|
14 |
end; |
|
15 |
||
16 |
fun negDivAlg (a,b) = |
|
17 |
if 0<=a+b then (~1,a+b) |
|
18 |
else let val (q,r) = negDivAlg(a, 2*b) |
|
19 |
in if 0<=r-b then (2*q+1, r-b) else (2*q, r) |
|
20 |
end; |
|
21 |
||
22 |
fun negateSnd (q,r:int) = (q,~r); |
|
23 |
||
24 |
fun divAlg (a,b) = if 0<=a then |
|
25 |
if b>0 then posDivAlg (a,b) |
|
26 |
else if a=0 then (0,0) |
|
27 |
else negateSnd (negDivAlg (~a,~b)) |
|
28 |
else |
|
29 |
if 0<b then negDivAlg (a,b) |
|
30 |
else negateSnd (posDivAlg (~a,~b)); |
|
6917 | 31 |
*) |
32 |
||
33 |
Addsimps [zless_nat_conj]; |
|
34 |
||
35 |
(*** Uniqueness and monotonicity of quotients and remainders ***) |
|
36 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
37 |
Goal "[| b*q' + r' <= b*q + r; #0 <= r'; #0 < b; r < b |] \ |
6917 | 38 |
\ ==> q' <= (q::int)"; |
39 |
by (subgoal_tac "r' + b * (q'-q) <= r" 1); |
|
40 |
by (simp_tac (simpset() addsimps zcompare_rls@[zdiff_zmult_distrib2]) 2); |
|
41 |
by (subgoal_tac "#0 < b * (#1 + q - q')" 1); |
|
42 |
by (etac order_le_less_trans 2); |
|
43 |
by (full_simp_tac (simpset() addsimps zcompare_rls@[zdiff_zmult_distrib2, |
|
44 |
zadd_zmult_distrib2]) 2); |
|
45 |
by (subgoal_tac "b * q' < b * (#1 + q)" 1); |
|
46 |
by (full_simp_tac (simpset() addsimps zcompare_rls@[zdiff_zmult_distrib2, |
|
47 |
zadd_zmult_distrib2]) 2); |
|
48 |
by (Asm_full_simp_tac 1); |
|
49 |
qed "unique_quotient_lemma"; |
|
50 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
51 |
Goal "[| b*q' + r' <= b*q + r; r <= #0; b < #0; b < r' |] \ |
6917 | 52 |
\ ==> q <= (q'::int)"; |
53 |
by (res_inst_tac [("b", "-b"), ("r", "-r'"), ("r'", "-r")] |
|
54 |
unique_quotient_lemma 1); |
|
55 |
by (auto_tac (claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
56 |
simpset() addsimps zcompare_rls@ |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
57 |
[zmult_zminus, zmult_zminus_right])); |
6917 | 58 |
qed "unique_quotient_lemma_neg"; |
59 |
||
60 |
||
61 |
Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \ |
|
62 |
\ ==> q = q'"; |
|
63 |
by (asm_full_simp_tac |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
64 |
(simpset() addsimps split_ifs@ |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
65 |
[quorem_def, linorder_neq_iff]) 1); |
6917 | 66 |
by Safe_tac; |
67 |
by (ALLGOALS Asm_full_simp_tac); |
|
68 |
by (REPEAT |
|
69 |
(blast_tac (claset() addIs [order_antisym] |
|
70 |
addDs [order_eq_refl RS unique_quotient_lemma, |
|
71 |
order_eq_refl RS unique_quotient_lemma_neg, |
|
72 |
sym]) 1)); |
|
73 |
qed "unique_quotient"; |
|
74 |
||
75 |
||
76 |
Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \ |
|
77 |
\ ==> r = r'"; |
|
78 |
by (subgoal_tac "q = q'" 1); |
|
79 |
by (blast_tac (claset() addIs [unique_quotient]) 2); |
|
80 |
by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1); |
|
81 |
qed "unique_remainder"; |
|
82 |
||
83 |
||
84 |
(*** Correctness of posDivAlg, the division algorithm for a>=0 and b>0 ***) |
|
85 |
||
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
86 |
(*Unfold all "let"s involving constants*) |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
87 |
Addsimps [read_instantiate_sg (sign_of IntDiv.thy) |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
88 |
[("s", "number_of ?v")] Let_def]; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
89 |
|
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
90 |
|
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
91 |
Goal "adjust a b (q,r) = (let diff = r-b in \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
92 |
\ if #0 <= diff then (#2*q + #1, diff) \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
93 |
\ else (#2*q, r))"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
94 |
by (simp_tac (simpset() addsimps [Let_def,adjust_def]) 1); |
6917 | 95 |
qed "adjust_eq"; |
96 |
Addsimps [adjust_eq]; |
|
97 |
||
98 |
(*Proving posDivAlg's termination condition*) |
|
99 |
val [tc] = posDivAlg.tcs; |
|
100 |
goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc)); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
101 |
by (auto_tac (claset(), simpset() addsimps [zmult_2])); |
6917 | 102 |
val lemma = result(); |
103 |
||
104 |
(* removing the termination condition from the generated theorems *) |
|
105 |
||
106 |
bind_thm ("posDivAlg_raw_eqn", lemma RS hd posDivAlg.rules); |
|
107 |
||
108 |
(**use with simproc to avoid re-proving the premise*) |
|
109 |
Goal "#0 < b ==> \ |
|
110 |
\ posDivAlg (a,b) = \ |
|
111 |
\ (if a<b then (#0,a) else adjust a b (posDivAlg(a, #2*b)))"; |
|
112 |
by (rtac (posDivAlg_raw_eqn RS trans) 1); |
|
113 |
by (Asm_simp_tac 1); |
|
114 |
qed "posDivAlg_eqn"; |
|
115 |
||
116 |
val posDivAlg_induct = lemma RS posDivAlg.induct; |
|
117 |
||
118 |
||
119 |
(*Correctness of posDivAlg: it computes quotients correctly*) |
|
120 |
Goal "#0 <= a --> #0 < b --> quorem ((a, b), posDivAlg (a, b))"; |
|
121 |
by (res_inst_tac [("u", "a"), ("v", "b")] posDivAlg_induct 1); |
|
122 |
by Auto_tac; |
|
123 |
by (ALLGOALS |
|
124 |
(asm_full_simp_tac (simpset() addsimps [quorem_def, |
|
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
125 |
pos_imp_zmult_pos_iff]))); |
6917 | 126 |
(*base case: a<b*) |
127 |
by (asm_full_simp_tac (simpset() addsimps [posDivAlg_eqn]) 1); |
|
128 |
(*main argument*) |
|
129 |
by (stac posDivAlg_eqn 1); |
|
130 |
by (ALLGOALS Asm_simp_tac); |
|
131 |
by (etac splitE 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
132 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
133 |
simpset() addsimps zmult_ac@[zadd_zmult_distrib2, Let_def])); |
6917 | 134 |
qed_spec_mp "posDivAlg_correct"; |
135 |
||
136 |
||
137 |
(*** Correctness of negDivAlg, the division algorithm for a<0 and b>0 ***) |
|
138 |
||
139 |
(*Proving negDivAlg's termination condition*) |
|
140 |
val [tc] = negDivAlg.tcs; |
|
141 |
goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc)); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
142 |
by (auto_tac (claset(), simpset() addsimps [zmult_2])); |
6917 | 143 |
val lemma = result(); |
144 |
||
145 |
(* removing the termination condition from the generated theorems *) |
|
146 |
||
147 |
bind_thm ("negDivAlg_raw_eqn", lemma RS hd negDivAlg.rules); |
|
148 |
||
149 |
(**use with simproc to avoid re-proving the premise*) |
|
150 |
Goal "#0 < b ==> \ |
|
151 |
\ negDivAlg (a,b) = \ |
|
152 |
\ (if #0<=a+b then (#-1,a+b) else adjust a b (negDivAlg(a, #2*b)))"; |
|
153 |
by (rtac (negDivAlg_raw_eqn RS trans) 1); |
|
154 |
by (Asm_simp_tac 1); |
|
155 |
qed "negDivAlg_eqn"; |
|
156 |
||
157 |
val negDivAlg_induct = lemma RS negDivAlg.induct; |
|
158 |
||
159 |
||
160 |
(*Correctness of negDivAlg: it computes quotients correctly |
|
161 |
It doesn't work if a=0 because the 0/b=0 rather than -1*) |
|
162 |
Goal "a < #0 --> #0 < b --> quorem ((a, b), negDivAlg (a, b))"; |
|
163 |
by (res_inst_tac [("u", "a"), ("v", "b")] negDivAlg_induct 1); |
|
164 |
by Auto_tac; |
|
165 |
by (ALLGOALS |
|
166 |
(asm_full_simp_tac (simpset() addsimps [quorem_def, |
|
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
167 |
pos_imp_zmult_pos_iff]))); |
6917 | 168 |
(*base case: 0<=a+b*) |
169 |
by (asm_full_simp_tac (simpset() addsimps [negDivAlg_eqn]) 1); |
|
170 |
(*main argument*) |
|
171 |
by (stac negDivAlg_eqn 1); |
|
172 |
by (ALLGOALS Asm_simp_tac); |
|
173 |
by (etac splitE 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
174 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
175 |
simpset() addsimps zmult_ac@[zadd_zmult_distrib2, Let_def])); |
6917 | 176 |
qed_spec_mp "negDivAlg_correct"; |
177 |
||
178 |
||
179 |
(*** Existence shown by proving the division algorithm to be correct ***) |
|
180 |
||
181 |
(*the case a=0*) |
|
182 |
Goal "b ~= #0 ==> quorem ((#0,b), (#0,#0))"; |
|
183 |
by (auto_tac (claset(), |
|
184 |
simpset() addsimps [quorem_def, linorder_neq_iff])); |
|
185 |
qed "quorem_0"; |
|
186 |
||
187 |
Goal "posDivAlg (#0, b) = (#0, #0)"; |
|
188 |
by (stac posDivAlg_raw_eqn 1); |
|
189 |
by Auto_tac; |
|
190 |
qed "posDivAlg_0"; |
|
191 |
Addsimps [posDivAlg_0]; |
|
192 |
||
193 |
Goal "negDivAlg (#-1, b) = (#-1, b-#1)"; |
|
194 |
by (stac negDivAlg_raw_eqn 1); |
|
195 |
by Auto_tac; |
|
196 |
qed "negDivAlg_minus1"; |
|
197 |
Addsimps [negDivAlg_minus1]; |
|
198 |
||
199 |
Goalw [negateSnd_def] "negateSnd(q,r) = (q,-r)"; |
|
200 |
by Auto_tac; |
|
201 |
qed "negateSnd_eq"; |
|
202 |
Addsimps [negateSnd_eq]; |
|
203 |
||
204 |
Goal "quorem ((-a,-b), qr) ==> quorem ((a,b), negateSnd qr)"; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
205 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
206 |
simpset() addsimps split_ifs@[zmult_zminus, quorem_def])); |
6917 | 207 |
qed "quorem_neg"; |
208 |
||
209 |
Goal "b ~= #0 ==> quorem ((a,b), divAlg(a,b))"; |
|
210 |
by (auto_tac (claset(), |
|
211 |
simpset() addsimps [quorem_0, divAlg_def])); |
|
212 |
by (REPEAT_FIRST (resolve_tac [quorem_neg, posDivAlg_correct, |
|
213 |
negDivAlg_correct])); |
|
214 |
by (auto_tac (claset(), |
|
215 |
simpset() addsimps [quorem_def, linorder_neq_iff])); |
|
216 |
qed "divAlg_correct"; |
|
217 |
||
6992 | 218 |
(** Aribtrary definitions for division by zero. Useful to simplify |
219 |
certain equations **) |
|
220 |
||
221 |
Goal "a div (#0::int) = #0"; |
|
222 |
by (simp_tac (simpset() addsimps [div_def, divAlg_def, posDivAlg_raw_eqn]) 1); |
|
7035 | 223 |
qed "DIVISION_BY_ZERO_ZDIV"; (*NOT for adding to default simpset*) |
6992 | 224 |
|
225 |
Goal "a mod (#0::int) = a"; |
|
226 |
by (simp_tac (simpset() addsimps [mod_def, divAlg_def, posDivAlg_raw_eqn]) 1); |
|
7035 | 227 |
qed "DIVISION_BY_ZERO_ZMOD"; (*NOT for adding to default simpset*) |
6992 | 228 |
|
7035 | 229 |
fun zdiv_undefined_case_tac s i = |
6992 | 230 |
case_tac s i THEN |
7035 | 231 |
asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO_ZDIV, |
232 |
DIVISION_BY_ZERO_ZMOD]) i; |
|
6992 | 233 |
|
234 |
(** Basic laws about division and remainder **) |
|
235 |
||
236 |
Goal "(a::int) = b * (a div b) + (a mod b)"; |
|
7035 | 237 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6917 | 238 |
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1); |
239 |
by (auto_tac (claset(), |
|
240 |
simpset() addsimps [quorem_def, div_def, mod_def])); |
|
241 |
qed "zmod_zdiv_equality"; |
|
242 |
||
243 |
Goal "(#0::int) < b ==> #0 <= a mod b & a mod b < b"; |
|
244 |
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1); |
|
245 |
by (auto_tac (claset(), |
|
246 |
simpset() addsimps [quorem_def, mod_def])); |
|
247 |
bind_thm ("pos_mod_sign", result() RS conjunct1); |
|
248 |
bind_thm ("pos_mod_bound", result() RS conjunct2); |
|
249 |
||
250 |
Goal "b < (#0::int) ==> a mod b <= #0 & b < a mod b"; |
|
251 |
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1); |
|
252 |
by (auto_tac (claset(), |
|
253 |
simpset() addsimps [quorem_def, div_def, mod_def])); |
|
254 |
bind_thm ("neg_mod_sign", result() RS conjunct1); |
|
255 |
bind_thm ("neg_mod_bound", result() RS conjunct2); |
|
256 |
||
257 |
||
6992 | 258 |
(** proving general properties of div and mod **) |
259 |
||
260 |
Goal "b ~= #0 ==> quorem ((a, b), (a div b, a mod b))"; |
|
261 |
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1); |
|
262 |
by (auto_tac |
|
263 |
(claset(), |
|
264 |
simpset() addsimps [quorem_def, linorder_neq_iff, |
|
265 |
pos_mod_sign,pos_mod_bound, |
|
266 |
neg_mod_sign, neg_mod_bound])); |
|
267 |
qed "quorem_div_mod"; |
|
268 |
||
269 |
Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a div b = q"; |
|
270 |
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_quotient]) 1); |
|
271 |
qed "quorem_div"; |
|
272 |
||
273 |
Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a mod b = r"; |
|
274 |
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_remainder]) 1); |
|
275 |
qed "quorem_mod"; |
|
276 |
||
277 |
Goal "[| (#0::int) <= a; a < b |] ==> a div b = #0"; |
|
278 |
by (rtac quorem_div 1); |
|
279 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
|
6999 | 280 |
qed "div_pos_pos_trivial"; |
6992 | 281 |
|
282 |
Goal "[| a <= (#0::int); b < a |] ==> a div b = #0"; |
|
283 |
by (rtac quorem_div 1); |
|
284 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
|
6999 | 285 |
qed "div_neg_neg_trivial"; |
286 |
||
287 |
Goal "[| (#0::int) < a; a+b <= #0 |] ==> a div b = #-1"; |
|
288 |
by (rtac quorem_div 1); |
|
289 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
|
290 |
qed "div_pos_neg_trivial"; |
|
291 |
||
292 |
(*There is no div_neg_pos_trivial because #0 div b = #0 would supersede it*) |
|
6992 | 293 |
|
294 |
Goal "[| (#0::int) <= a; a < b |] ==> a mod b = a"; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
295 |
by (res_inst_tac [("q","#0")] quorem_mod 1); |
6992 | 296 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
6999 | 297 |
qed "mod_pos_pos_trivial"; |
6992 | 298 |
|
299 |
Goal "[| a <= (#0::int); b < a |] ==> a mod b = a"; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
300 |
by (res_inst_tac [("q","#0")] quorem_mod 1); |
6992 | 301 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
6999 | 302 |
qed "mod_neg_neg_trivial"; |
303 |
||
304 |
Goal "[| (#0::int) < a; a+b <= #0 |] ==> a mod b = a+b"; |
|
305 |
by (res_inst_tac [("q","#-1")] quorem_mod 1); |
|
306 |
by (auto_tac (claset(), simpset() addsimps [quorem_def])); |
|
307 |
qed "mod_pos_neg_trivial"; |
|
308 |
||
309 |
(*There is no mod_neg_pos_trivial...*) |
|
6992 | 310 |
|
311 |
||
312 |
(*Simpler laws such as -a div b = -(a div b) FAIL*) |
|
313 |
Goal "(-a) div (-b) = a div (b::int)"; |
|
7035 | 314 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 315 |
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) |
316 |
RS quorem_div) 1); |
|
317 |
by Auto_tac; |
|
318 |
qed "zdiv_zminus_zminus"; |
|
319 |
Addsimps [zdiv_zminus_zminus]; |
|
320 |
||
321 |
(*Simpler laws such as -a mod b = -(a mod b) FAIL*) |
|
322 |
Goal "(-a) mod (-b) = - (a mod (b::int))"; |
|
7035 | 323 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 324 |
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) |
325 |
RS quorem_mod) 1); |
|
326 |
by Auto_tac; |
|
327 |
qed "zmod_zminus_zminus"; |
|
328 |
Addsimps [zmod_zminus_zminus]; |
|
329 |
||
330 |
||
331 |
(*** division of a number by itself ***) |
|
332 |
||
333 |
Goal "[| (#0::int) < a; a = r + a*q; r < a |] ==> #1 <= q"; |
|
334 |
by (subgoal_tac "#0 < a*q" 1); |
|
335 |
by (arith_tac 2); |
|
336 |
by (asm_full_simp_tac (simpset() addsimps [pos_imp_zmult_pos_iff]) 1); |
|
337 |
val lemma1 = result(); |
|
338 |
||
339 |
Goal "[| (#0::int) < a; a = r + a*q; #0 <= r |] ==> q <= #1"; |
|
340 |
by (subgoal_tac "#0 <= a*(#1-q)" 1); |
|
341 |
by (asm_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2); |
|
342 |
by (asm_full_simp_tac (simpset() addsimps [pos_imp_zmult_nonneg_iff]) 1); |
|
343 |
val lemma2 = result(); |
|
344 |
||
345 |
Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> q = #1"; |
|
346 |
by (asm_full_simp_tac |
|
347 |
(simpset() addsimps split_ifs@[quorem_def, linorder_neq_iff]) 1); |
|
348 |
by (rtac order_antisym 1); |
|
349 |
by Safe_tac; |
|
350 |
by Auto_tac; |
|
351 |
by (res_inst_tac [("a", "-a"),("r", "-r")] lemma1 3); |
|
352 |
by (res_inst_tac [("a", "-a"),("r", "-r")] lemma2 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
353 |
by (REPEAT (force_tac (claset() addIs [lemma1,lemma2], |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
354 |
simpset() addsimps [zadd_commute, zmult_zminus]) 1)); |
6992 | 355 |
qed "self_quotient"; |
356 |
||
357 |
Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> r = #0"; |
|
7499 | 358 |
by (ftac self_quotient 1); |
6992 | 359 |
by (assume_tac 1); |
360 |
by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1); |
|
361 |
qed "self_remainder"; |
|
362 |
||
363 |
Goal "a ~= #0 ==> a div a = (#1::int)"; |
|
364 |
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_quotient]) 1); |
|
365 |
qed "zdiv_self"; |
|
366 |
Addsimps [zdiv_self]; |
|
367 |
||
368 |
(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) |
|
369 |
Goal "a mod a = (#0::int)"; |
|
7035 | 370 |
by (zdiv_undefined_case_tac "a = #0" 1); |
6992 | 371 |
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_remainder]) 1); |
372 |
qed "zmod_self"; |
|
373 |
Addsimps [zmod_self]; |
|
374 |
||
375 |
||
6917 | 376 |
(*** Computation of division and remainder ***) |
377 |
||
378 |
Goal "(#0::int) div b = #0"; |
|
379 |
by (simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
380 |
qed "zdiv_zero"; |
6917 | 381 |
|
7035 | 382 |
Goal "(#0::int) < b ==> #-1 div b = #-1"; |
383 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
|
384 |
qed "div_eq_minus1"; |
|
385 |
||
6917 | 386 |
Goal "(#0::int) mod b = #0"; |
387 |
by (simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
388 |
qed "zmod_zero"; |
6917 | 389 |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
390 |
Addsimps [zdiv_zero, zmod_zero]; |
6917 | 391 |
|
7035 | 392 |
Goal "(#0::int) < b ==> #-1 div b = #-1"; |
393 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
394 |
qed "zdiv_minus1"; |
7035 | 395 |
|
396 |
Goal "(#0::int) < b ==> #-1 mod b = b-#1"; |
|
397 |
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
398 |
qed "zmod_minus1"; |
7035 | 399 |
|
6917 | 400 |
(** a positive, b positive **) |
401 |
||
6992 | 402 |
Goal "[| #0 < a; #0 <= b |] ==> a div b = fst (posDivAlg(a,b))"; |
6917 | 403 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
404 |
qed "div_pos_pos"; |
|
405 |
||
6992 | 406 |
Goal "[| #0 < a; #0 <= b |] ==> a mod b = snd (posDivAlg(a,b))"; |
6917 | 407 |
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
408 |
qed "mod_pos_pos"; |
|
409 |
||
410 |
(** a negative, b positive **) |
|
411 |
||
412 |
Goal "[| a < #0; #0 < b |] ==> a div b = fst (negDivAlg(a,b))"; |
|
413 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
|
414 |
qed "div_neg_pos"; |
|
415 |
||
416 |
Goal "[| a < #0; #0 < b |] ==> a mod b = snd (negDivAlg(a,b))"; |
|
417 |
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
|
418 |
qed "mod_neg_pos"; |
|
419 |
||
420 |
(** a positive, b negative **) |
|
421 |
||
422 |
Goal "[| #0 < a; b < #0 |] ==> a div b = fst (negateSnd(negDivAlg(-a,-b)))"; |
|
423 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
|
424 |
qed "div_pos_neg"; |
|
425 |
||
426 |
Goal "[| #0 < a; b < #0 |] ==> a mod b = snd (negateSnd(negDivAlg(-a,-b)))"; |
|
427 |
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
|
428 |
qed "mod_pos_neg"; |
|
429 |
||
430 |
(** a negative, b negative **) |
|
431 |
||
6992 | 432 |
Goal "[| a < #0; b <= #0 |] ==> a div b = fst (negateSnd(posDivAlg(-a,-b)))"; |
6917 | 433 |
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1); |
434 |
qed "div_neg_neg"; |
|
435 |
||
6992 | 436 |
Goal "[| a < #0; b <= #0 |] ==> a mod b = snd (negateSnd(posDivAlg(-a,-b)))"; |
6917 | 437 |
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1); |
438 |
qed "mod_neg_neg"; |
|
439 |
||
440 |
Addsimps (map (read_instantiate_sg (sign_of IntDiv.thy) |
|
441 |
[("a", "number_of ?v"), ("b", "number_of ?w")]) |
|
442 |
[div_pos_pos, div_neg_pos, div_pos_neg, div_neg_neg, |
|
443 |
mod_pos_pos, mod_neg_pos, mod_pos_neg, mod_neg_neg, |
|
444 |
posDivAlg_eqn, negDivAlg_eqn]); |
|
6992 | 445 |
|
6917 | 446 |
|
447 |
(** Special-case simplification **) |
|
448 |
||
449 |
Goal "a mod (#1::int) = #0"; |
|
450 |
by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_sign 1); |
|
451 |
by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_bound 2); |
|
452 |
by Auto_tac; |
|
453 |
qed "zmod_1"; |
|
454 |
Addsimps [zmod_1]; |
|
455 |
||
456 |
Goal "a div (#1::int) = a"; |
|
457 |
by (cut_inst_tac [("a","a"),("b","#1")] zmod_zdiv_equality 1); |
|
458 |
by Auto_tac; |
|
459 |
qed "zdiv_1"; |
|
460 |
Addsimps [zdiv_1]; |
|
461 |
||
462 |
Goal "a mod (#-1::int) = #0"; |
|
463 |
by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_sign 1); |
|
464 |
by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_bound 2); |
|
465 |
by Auto_tac; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
466 |
qed "zmod_minus1_right"; |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
467 |
Addsimps [zmod_minus1_right]; |
6917 | 468 |
|
469 |
Goal "a div (#-1::int) = -a"; |
|
470 |
by (cut_inst_tac [("a","a"),("b","#-1")] zmod_zdiv_equality 1); |
|
471 |
by Auto_tac; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
472 |
qed "zdiv_minus1_right"; |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
473 |
Addsimps [zdiv_minus1_right]; |
6917 | 474 |
|
475 |
||
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
476 |
(*** Monotonicity in the first argument (divisor) ***) |
6917 | 477 |
|
478 |
Goal "[| a <= a'; #0 < (b::int) |] ==> a div b <= a' div b"; |
|
479 |
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1); |
|
6992 | 480 |
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1); |
6917 | 481 |
by (rtac unique_quotient_lemma 1); |
482 |
by (etac subst 1); |
|
483 |
by (etac subst 1); |
|
484 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); |
|
485 |
qed "zdiv_mono1"; |
|
486 |
||
487 |
Goal "[| a <= a'; (b::int) < #0 |] ==> a' div b <= a div b"; |
|
488 |
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1); |
|
6992 | 489 |
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1); |
6917 | 490 |
by (rtac unique_quotient_lemma_neg 1); |
491 |
by (etac subst 1); |
|
492 |
by (etac subst 1); |
|
493 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [neg_mod_sign,neg_mod_bound]))); |
|
494 |
qed "zdiv_mono1_neg"; |
|
495 |
||
496 |
||
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
497 |
(*** Monotonicity in the second argument (dividend) ***) |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
498 |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
499 |
Goal "[| b*q + r = b'*q' + r'; #0 <= b'*q' + r'; \ |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
500 |
\ r' < b'; #0 <= r; #0 < b'; b' <= b |] \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
501 |
\ ==> q <= (q'::int)"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
502 |
by (subgoal_tac "#0 <= q'" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
503 |
by (subgoal_tac "#0 < b'*(q' + #1)" 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
504 |
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 3); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
505 |
by (asm_full_simp_tac (simpset() addsimps [pos_imp_zmult_pos_iff]) 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
506 |
by (subgoal_tac "b*q < b*(q' + #1)" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
507 |
by (Asm_full_simp_tac 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
508 |
by (subgoal_tac "b*q = r' - r + b'*q'" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
509 |
by (simp_tac (simpset() addsimps zcompare_rls) 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
510 |
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
511 |
by (stac zadd_commute 1 THEN rtac zadd_zless_mono 1 THEN arith_tac 1); |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
512 |
by (rtac zmult_zle_mono1 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
513 |
by Auto_tac; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
514 |
qed "zdiv_mono2_lemma"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
515 |
|
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
516 |
Goal "[| (#0::int) <= a; #0 < b'; b' <= b |] \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
517 |
\ ==> a div b <= a div b'"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
518 |
by (subgoal_tac "b ~= #0" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
519 |
by (arith_tac 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
520 |
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1); |
6992 | 521 |
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1); |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
522 |
by (rtac zdiv_mono2_lemma 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
523 |
by (etac subst 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
524 |
by (etac subst 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
525 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
526 |
qed "zdiv_mono2"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
527 |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
528 |
Goal "[| b*q + r = b'*q' + r'; b'*q' + r' < #0; \ |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
529 |
\ r < b; #0 <= r'; #0 < b'; b' <= b |] \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
530 |
\ ==> q' <= (q::int)"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
531 |
by (subgoal_tac "q' < #0" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
532 |
by (subgoal_tac "b'*q' < #0" 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
533 |
by (arith_tac 3); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
534 |
by (asm_full_simp_tac (simpset() addsimps [pos_imp_zmult_neg_iff]) 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
535 |
by (subgoal_tac "b*q' < b*(q + #1)" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
536 |
by (Asm_full_simp_tac 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
537 |
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
538 |
by (subgoal_tac "b*q' <= b'*q'" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
539 |
by (asm_simp_tac (simpset() addsimps [zmult_zle_mono1_neg]) 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
540 |
by (subgoal_tac "b'*q' < b + b*q" 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
541 |
by (Asm_simp_tac 2); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
542 |
by (arith_tac 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
543 |
qed "zdiv_mono2_neg_lemma"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
544 |
|
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
545 |
Goal "[| a < (#0::int); #0 < b'; b' <= b |] \ |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
546 |
\ ==> a div b' <= a div b"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
547 |
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1); |
6992 | 548 |
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1); |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
549 |
by (rtac zdiv_mono2_neg_lemma 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
550 |
by (etac subst 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
551 |
by (etac subst 1); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
552 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound]))); |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
553 |
qed "zdiv_mono2_neg"; |
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
554 |
|
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
555 |
|
6992 | 556 |
(*** More algebraic laws for div and mod ***) |
6943
2cde117d2738
faster division algorithm; monotonicity of div in 2nd arg
paulson
parents:
6917
diff
changeset
|
557 |
|
6992 | 558 |
(** proving (a*b) div c = a * (b div c) + a * (b mod c) **) |
559 |
||
560 |
Goal "[| quorem((b,c),(q,r)); c ~= #0 |] \ |
|
561 |
\ ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))"; |
|
562 |
by (auto_tac |
|
563 |
(claset(), |
|
7549 | 564 |
simpset() delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
565 |
addsimps split_ifs@zmult_ac@ |
|
6992 | 566 |
[quorem_def, linorder_neq_iff, |
567 |
zadd_zmult_distrib2, |
|
568 |
pos_mod_sign,pos_mod_bound, |
|
569 |
neg_mod_sign, neg_mod_bound])); |
|
7549 | 570 |
by (ALLGOALS(rtac zmod_zdiv_equality)); |
6992 | 571 |
val lemma = result(); |
572 |
||
573 |
Goal "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"; |
|
7035 | 574 |
by (zdiv_undefined_case_tac "c = #0" 1); |
6992 | 575 |
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_div]) 1); |
576 |
qed "zdiv_zmult1_eq"; |
|
577 |
||
578 |
Goal "(a*b) mod c = a*(b mod c) mod (c::int)"; |
|
7035 | 579 |
by (zdiv_undefined_case_tac "c = #0" 1); |
6992 | 580 |
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_mod]) 1); |
581 |
qed "zmod_zmult1_eq"; |
|
582 |
||
583 |
Goal "b ~= (#0::int) ==> (a*b) div b = a"; |
|
584 |
by (asm_simp_tac (simpset() addsimps [zdiv_zmult1_eq]) 1); |
|
585 |
qed "zdiv_zmult_self1"; |
|
586 |
||
587 |
Goal "b ~= (#0::int) ==> (b*a) div b = a"; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
588 |
by (stac zmult_commute 1 THEN etac zdiv_zmult_self1 1); |
6992 | 589 |
qed "zdiv_zmult_self2"; |
590 |
||
591 |
Addsimps [zdiv_zmult_self1, zdiv_zmult_self2]; |
|
592 |
||
593 |
Goal "(a*b) mod b = (#0::int)"; |
|
594 |
by (simp_tac (simpset() addsimps [zmod_zmult1_eq]) 1); |
|
595 |
qed "zmod_zmult_self1"; |
|
596 |
||
597 |
Goal "(b*a) mod b = (#0::int)"; |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
598 |
by (simp_tac (simpset() addsimps [zmult_commute, zmod_zmult1_eq]) 1); |
6992 | 599 |
qed "zmod_zmult_self2"; |
600 |
||
601 |
Addsimps [zmod_zmult_self1, zmod_zmult_self2]; |
|
602 |
||
603 |
||
604 |
(** proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) **) |
|
6917 | 605 |
|
6992 | 606 |
Goal "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); c ~= #0 |] \ |
607 |
\ ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))"; |
|
608 |
by (auto_tac |
|
609 |
(claset(), |
|
7549 | 610 |
simpset() delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
611 |
addsimps split_ifs@zmult_ac@ |
|
6992 | 612 |
[quorem_def, linorder_neq_iff, |
613 |
zadd_zmult_distrib2, |
|
614 |
pos_mod_sign,pos_mod_bound, |
|
615 |
neg_mod_sign, neg_mod_bound])); |
|
7549 | 616 |
by (ALLGOALS(rtac zmod_zdiv_equality)); |
6992 | 617 |
val lemma = result(); |
618 |
||
6999 | 619 |
(*NOT suitable for rewriting: the RHS has an instance of the LHS*) |
6992 | 620 |
Goal "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"; |
7035 | 621 |
by (zdiv_undefined_case_tac "c = #0" 1); |
6992 | 622 |
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod] |
623 |
MRS lemma RS quorem_div]) 1); |
|
624 |
qed "zdiv_zadd1_eq"; |
|
625 |
||
626 |
Goal "(a+b) mod (c::int) = (a mod c + b mod c) mod c"; |
|
7035 | 627 |
by (zdiv_undefined_case_tac "c = #0" 1); |
6992 | 628 |
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod] |
629 |
MRS lemma RS quorem_mod]) 1); |
|
630 |
qed "zmod_zadd1_eq"; |
|
631 |
||
632 |
||
633 |
Goal "(a mod b) div b = (#0::int)"; |
|
7035 | 634 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 635 |
by (auto_tac (claset(), |
6999 | 636 |
simpset() addsimps [linorder_neq_iff, |
637 |
pos_mod_sign, pos_mod_bound, div_pos_pos_trivial, |
|
638 |
neg_mod_sign, neg_mod_bound, div_neg_neg_trivial])); |
|
6992 | 639 |
qed "mod_div_trivial"; |
640 |
Addsimps [mod_div_trivial]; |
|
641 |
||
642 |
Goal "(a mod b) mod b = a mod (b::int)"; |
|
7035 | 643 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 644 |
by (auto_tac (claset(), |
6999 | 645 |
simpset() addsimps [linorder_neq_iff, |
646 |
pos_mod_sign, pos_mod_bound, mod_pos_pos_trivial, |
|
647 |
neg_mod_sign, neg_mod_bound, mod_neg_neg_trivial])); |
|
6992 | 648 |
qed "mod_mod_trivial"; |
649 |
Addsimps [mod_mod_trivial]; |
|
650 |
||
651 |
||
652 |
Goal "a ~= (#0::int) ==> (a+b) div a = b div a + #1"; |
|
653 |
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1); |
|
654 |
qed "zdiv_zadd_self1"; |
|
655 |
||
656 |
Goal "a ~= (#0::int) ==> (b+a) div a = b div a + #1"; |
|
657 |
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1); |
|
658 |
qed "zdiv_zadd_self2"; |
|
659 |
Addsimps [zdiv_zadd_self1, zdiv_zadd_self2]; |
|
660 |
||
661 |
Goal "(a+b) mod a = b mod (a::int)"; |
|
7035 | 662 |
by (zdiv_undefined_case_tac "a = #0" 1); |
6992 | 663 |
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
664 |
qed "zmod_zadd_self1"; |
|
665 |
||
666 |
Goal "(b+a) mod a = b mod (a::int)"; |
|
7035 | 667 |
by (zdiv_undefined_case_tac "a = #0" 1); |
6992 | 668 |
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); |
669 |
qed "zmod_zadd_self2"; |
|
670 |
Addsimps [zmod_zadd_self1, zmod_zadd_self2]; |
|
671 |
||
672 |
||
673 |
(*** proving a div (b*c) = (a div b) div c ***) |
|
674 |
||
675 |
(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but |
|
676 |
7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems |
|
677 |
to cause particular problems.*) |
|
678 |
||
679 |
(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) |
|
6917 | 680 |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
681 |
Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b*c < b*(q mod c) + r"; |
6992 | 682 |
by (subgoal_tac "b * (c - q mod c) < r * #1" 1); |
683 |
by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1); |
|
684 |
by (rtac order_le_less_trans 1); |
|
685 |
by (etac zmult_zless_mono1 2); |
|
686 |
by (rtac zmult_zle_mono2_neg 1); |
|
687 |
by (auto_tac |
|
688 |
(claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
689 |
simpset() addsimps zcompare_rls@ |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
690 |
[zadd_commute, add1_zle_eq, pos_mod_bound])); |
6992 | 691 |
val lemma1 = result(); |
692 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
693 |
Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b * (q mod c) + r <= #0"; |
6992 | 694 |
by (subgoal_tac "b * (q mod c) <= #0" 1); |
695 |
by (arith_tac 1); |
|
696 |
by (asm_simp_tac (simpset() addsimps [neg_imp_zmult_nonpos_iff, |
|
697 |
pos_mod_sign]) 1); |
|
698 |
val lemma2 = result(); |
|
699 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
700 |
Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> #0 <= b * (q mod c) + r"; |
6992 | 701 |
by (subgoal_tac "#0 <= b * (q mod c)" 1); |
702 |
by (arith_tac 1); |
|
703 |
by (asm_simp_tac |
|
704 |
(simpset() addsimps [pos_imp_zmult_nonneg_iff, pos_mod_sign]) 1); |
|
705 |
val lemma3 = result(); |
|
706 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
707 |
Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> b * (q mod c) + r < b * c"; |
6992 | 708 |
by (subgoal_tac "r * #1 < b * (c - q mod c)" 1); |
709 |
by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1); |
|
710 |
by (rtac order_less_le_trans 1); |
|
711 |
by (etac zmult_zless_mono1 1); |
|
712 |
by (rtac zmult_zle_mono2 2); |
|
713 |
by (auto_tac |
|
714 |
(claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
715 |
simpset() addsimps zcompare_rls@ |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
716 |
[zadd_commute, add1_zle_eq, pos_mod_bound])); |
6992 | 717 |
val lemma4 = result(); |
718 |
||
719 |
Goal "[| quorem ((a,b), (q,r)); b ~= #0; #0 < c |] \ |
|
720 |
\ ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))"; |
|
721 |
by (auto_tac (*SLOW*) |
|
722 |
(claset(), |
|
7549 | 723 |
simpset() delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
724 |
addsimps split_ifs@zmult_ac@ |
|
6992 | 725 |
[quorem_def, linorder_neq_iff, |
726 |
pos_imp_zmult_pos_iff, |
|
727 |
neg_imp_zmult_pos_iff, |
|
728 |
zadd_zmult_distrib2 RS sym, |
|
729 |
lemma1, lemma2, lemma3, lemma4])); |
|
7549 | 730 |
by (ALLGOALS(rtac zmod_zdiv_equality)); |
6992 | 731 |
val lemma = result(); |
732 |
||
733 |
Goal "(#0::int) < c ==> a div (b*c) = (a div b) div c"; |
|
7035 | 734 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 735 |
by (force_tac (claset(), |
736 |
simpset() addsimps [quorem_div_mod RS lemma RS quorem_div, |
|
737 |
zmult_eq_0_iff]) 1); |
|
738 |
qed "zdiv_zmult2_eq"; |
|
6917 | 739 |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
740 |
Goal "(#0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"; |
7035 | 741 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 742 |
by (force_tac (claset(), |
743 |
simpset() addsimps [quorem_div_mod RS lemma RS quorem_mod, |
|
744 |
zmult_eq_0_iff]) 1); |
|
745 |
qed "zmod_zmult2_eq"; |
|
746 |
||
747 |
||
748 |
(*** Cancellation of common factors in "div" ***) |
|
749 |
||
750 |
Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) div (c*b) = a div b"; |
|
751 |
by (stac zdiv_zmult2_eq 1); |
|
752 |
by Auto_tac; |
|
753 |
val lemma1 = result(); |
|
754 |
||
755 |
Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) div (c*b) = a div b"; |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
756 |
by (subgoal_tac "(c * (-a)) div (c * (-b)) = (-a) div (-b)" 1); |
6992 | 757 |
by (rtac lemma1 2); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
758 |
by (auto_tac (claset(), simpset() addsimps [zmult_zminus_right])); |
6992 | 759 |
val lemma2 = result(); |
760 |
||
761 |
Goal "c ~= (#0::int) ==> (c*a) div (c*b) = a div b"; |
|
7035 | 762 |
by (zdiv_undefined_case_tac "b = #0" 1); |
6992 | 763 |
by (auto_tac |
764 |
(claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
765 |
simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff, |
6992 | 766 |
lemma1, lemma2])); |
767 |
qed "zdiv_zmult_zmult1"; |
|
6917 | 768 |
|
6992 | 769 |
Goal "c ~= (#0::int) ==> (a*c) div (b*c) = a div b"; |
770 |
by (dtac zdiv_zmult_zmult1 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
771 |
by (auto_tac (claset(), simpset() addsimps [zmult_commute])); |
6992 | 772 |
qed "zdiv_zmult_zmult2"; |
773 |
||
774 |
||
775 |
||
776 |
(*** Distribution of factors over "mod" ***) |
|
777 |
||
778 |
Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)"; |
|
779 |
by (stac zmod_zmult2_eq 1); |
|
780 |
by Auto_tac; |
|
781 |
val lemma1 = result(); |
|
782 |
||
783 |
Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)"; |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
784 |
by (subgoal_tac "(c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))" 1); |
6992 | 785 |
by (rtac lemma1 2); |
786 |
by (auto_tac (claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
787 |
simpset() addsimps [zmult_zminus_right, zmod_zminus_zminus])); |
6992 | 788 |
val lemma2 = result(); |
789 |
||
6999 | 790 |
Goal "(c*a) mod (c*b) = (c::int) * (a mod b)"; |
7035 | 791 |
by (zdiv_undefined_case_tac "b = #0" 1); |
792 |
by (zdiv_undefined_case_tac "c = #0" 1); |
|
6992 | 793 |
by (auto_tac |
794 |
(claset(), |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
795 |
simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff, |
6992 | 796 |
lemma1, lemma2])); |
797 |
qed "zmod_zmult_zmult1"; |
|
798 |
||
6999 | 799 |
Goal "(a*c) mod (b*c) = (a mod b) * (c::int)"; |
800 |
by (cut_inst_tac [("c","c")] zmod_zmult_zmult1 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
801 |
by (auto_tac (claset(), simpset() addsimps [zmult_commute])); |
6992 | 802 |
qed "zmod_zmult_zmult2"; |
803 |
||
804 |
||
6999 | 805 |
(*** Speeding up the division algorithm with shifting ***) |
6992 | 806 |
|
7035 | 807 |
(** computing "div" by shifting **) |
6999 | 808 |
|
809 |
Goal "(#0::int) <= a ==> (#1 + #2*b) div (#2*a) = b div a"; |
|
7035 | 810 |
by (zdiv_undefined_case_tac "a = #0" 1); |
6999 | 811 |
by (subgoal_tac "#1 <= a" 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
812 |
by (arith_tac 2); |
6999 | 813 |
by (subgoal_tac "#1 < a * #2" 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
814 |
by (dres_inst_tac [("i","#1"), ("k", "#2")] zmult_zle_mono1 2); |
6999 | 815 |
by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
816 |
by (rtac zmult_zle_mono2 2); |
6999 | 817 |
by (auto_tac (claset(), |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
818 |
simpset() addsimps [zadd_commute, zmult_commute, |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
819 |
add1_zle_eq, pos_mod_bound])); |
6999 | 820 |
by (stac zdiv_zadd1_eq 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
821 |
by (asm_simp_tac (simpset() addsimps [zdiv_zmult_zmult2, zmod_zmult_zmult2, |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
822 |
div_pos_pos_trivial]) 1); |
6999 | 823 |
by (stac div_pos_pos_trivial 1); |
7549 | 824 |
by (asm_simp_tac (simpset() delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
825 |
addsimps zadd_ac@ [zmult_2_right, mod_pos_pos_trivial, |
|
826 |
pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1); |
|
6999 | 827 |
by (auto_tac (claset(), |
828 |
simpset() addsimps [mod_pos_pos_trivial])); |
|
829 |
qed "pos_zdiv_times_2"; |
|
830 |
||
831 |
||
832 |
Goal "a <= (#0::int) ==> (#1 + #2*b) div (#2*a) = (b+#1) div a"; |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
833 |
by (subgoal_tac "(#1 + #2*(-b-#1)) div (#2 * (-a)) = (-b-#1) div (-a)" 1); |
7035 | 834 |
by (rtac pos_zdiv_times_2 2); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
835 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
836 |
simpset() addsimps [zmult_zminus_right])); |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
837 |
by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1); |
6999 | 838 |
by (Simp_tac 2); |
839 |
by (asm_full_simp_tac (HOL_ss |
|
840 |
addsimps [zdiv_zminus_zminus, zdiff_def, |
|
841 |
zminus_zadd_distrib RS sym]) 1); |
|
842 |
qed "neg_zdiv_times_2"; |
|
843 |
||
844 |
||
845 |
(*Not clear why this must be proved separately; probably number_of causes |
|
846 |
simplification problems*) |
|
847 |
Goal "~ #0 <= x ==> x <= (#0::int)"; |
|
7035 | 848 |
by Auto_tac; |
6999 | 849 |
val lemma = result(); |
850 |
||
851 |
Goal "number_of (v BIT b) div number_of (w BIT False) = \ |
|
852 |
\ (if ~b | (#0::int) <= number_of w \ |
|
853 |
\ then number_of v div (number_of w) \ |
|
854 |
\ else (number_of v + (#1::int)) div (number_of w))"; |
|
855 |
by (simp_tac (simpset_of Int.thy |
|
856 |
addsimps [zadd_assoc, number_of_BIT]) 1); |
|
7035 | 857 |
by (asm_simp_tac (simpset() |
7549 | 858 |
delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
859 |
delsimps bin_arith_extra_simps@bin_rel_simps |
|
7035 | 860 |
addsimps [zmult_2 RS sym, zdiv_zmult_zmult1, |
861 |
pos_zdiv_times_2, lemma, neg_zdiv_times_2]) 1); |
|
6999 | 862 |
qed "zdiv_number_of_BIT"; |
863 |
||
864 |
Addsimps [zdiv_number_of_BIT]; |
|
865 |
||
866 |
||
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
867 |
(** computing "mod" by shifting (proofs resemble those for "div") **) |
7035 | 868 |
|
869 |
Goal "(#0::int) <= a ==> (#1 + #2*b) mod (#2*a) = #1 + #2 * (b mod a)"; |
|
870 |
by (zdiv_undefined_case_tac "a = #0" 1); |
|
871 |
by (subgoal_tac "#1 <= a" 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
872 |
by (arith_tac 2); |
7035 | 873 |
by (subgoal_tac "#1 < a * #2" 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
874 |
by (dres_inst_tac [("i","#1"), ("k", "#2")] zmult_zle_mono1 2); |
7035 | 875 |
by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
876 |
by (rtac zmult_zle_mono2 2); |
7035 | 877 |
by (auto_tac (claset(), |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
878 |
simpset() addsimps [zadd_commute, zmult_commute, |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
879 |
add1_zle_eq, pos_mod_bound])); |
7035 | 880 |
by (stac zmod_zadd1_eq 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
881 |
by (asm_simp_tac (simpset() addsimps [zmod_zmult_zmult2, |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
882 |
mod_pos_pos_trivial]) 1); |
7035 | 883 |
by (rtac mod_pos_pos_trivial 1); |
7549 | 884 |
by (asm_simp_tac (simpset() delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
885 |
addsimps zadd_ac@ [zmult_2_right, mod_pos_pos_trivial, |
|
886 |
pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1); |
|
7035 | 887 |
by (auto_tac (claset(), |
888 |
simpset() addsimps [mod_pos_pos_trivial])); |
|
889 |
qed "pos_zmod_times_2"; |
|
890 |
||
891 |
||
892 |
Goal "a <= (#0::int) ==> (#1 + #2*b) mod (#2*a) = #2 * ((b+#1) mod a) - #1"; |
|
893 |
by (subgoal_tac |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
894 |
"(#1 + #2*(-b-#1)) mod (#2*(-a)) = #1 + #2*((-b-#1) mod (-a))" 1); |
7035 | 895 |
by (rtac pos_zmod_times_2 2); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
896 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
897 |
simpset() addsimps [zmult_zminus_right])); |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
898 |
by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1); |
7035 | 899 |
by (Simp_tac 2); |
900 |
by (asm_full_simp_tac (HOL_ss |
|
901 |
addsimps [zmod_zminus_zminus, zdiff_def, |
|
902 |
zminus_zadd_distrib RS sym]) 1); |
|
7086 | 903 |
by (dtac (zminus_equation RS iffD1 RS sym) 1); |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
904 |
by (auto_tac (claset(), |
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
905 |
simpset() addsimps [zmult_zminus_right])); |
7035 | 906 |
qed "neg_zmod_times_2"; |
907 |
||
908 |
Goal "number_of (v BIT b) mod number_of (w BIT False) = \ |
|
909 |
\ (if b then \ |
|
910 |
\ if (#0::int) <= number_of w \ |
|
911 |
\ then #2 * (number_of v mod number_of w) + #1 \ |
|
912 |
\ else #2 * ((number_of v + (#1::int)) mod number_of w) - #1 \ |
|
913 |
\ else #2 * (number_of v mod number_of w))"; |
|
914 |
by (simp_tac (simpset_of Int.thy |
|
915 |
addsimps [zadd_assoc, number_of_BIT]) 1); |
|
916 |
by (asm_simp_tac (simpset() |
|
7549 | 917 |
delsimprocs [Int_CC.sum_conv, Int_CC.rel_conv] |
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
918 |
delsimps bin_arith_extra_simps@bin_rel_simps |
7035 | 919 |
addsimps [zmult_2 RS sym, zmod_zmult_zmult1, |
920 |
pos_zmod_times_2, lemma, neg_zmod_times_2]) 1); |
|
921 |
qed "zmod_number_of_BIT"; |
|
922 |
||
923 |
Addsimps [zmod_number_of_BIT]; |
|
924 |
||
925 |
||
926 |
(** Quotients of signs **) |
|
927 |
||
928 |
Goal "[| a < (#0::int); #0 < b |] ==> a div b < #0"; |
|
929 |
by (subgoal_tac "a div b <= #-1" 1); |
|
930 |
by (Force_tac 1); |
|
931 |
by (rtac order_trans 1); |
|
932 |
by (res_inst_tac [("a'","#-1")] zdiv_mono1 1); |
|
7074
e0730ffaafcc
zadd_ac and zmult_ac are no longer included by default
paulson
parents:
7035
diff
changeset
|
933 |
by (auto_tac (claset(), simpset() addsimps [zdiv_minus1])); |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
934 |
qed "div_neg_pos_less0"; |
7035 | 935 |
|
936 |
Goal "[| (#0::int) <= a; b < #0 |] ==> a div b <= #0"; |
|
937 |
by (dtac zdiv_mono1_neg 1); |
|
938 |
by Auto_tac; |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
939 |
qed "div_nonneg_neg_le0"; |
7035 | 940 |
|
941 |
Goal "(#0::int) < b ==> (#0 <= a div b) = (#0 <= a)"; |
|
942 |
by Auto_tac; |
|
943 |
by (dtac zdiv_mono1 2); |
|
944 |
by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff])); |
|
945 |
by (full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
7086
diff
changeset
|
946 |
by (blast_tac (claset() addIs [div_neg_pos_less0]) 1); |
7035 | 947 |
qed "pos_imp_zdiv_nonneg_iff"; |
948 |
||
949 |
Goal "b < (#0::int) ==> (#0 <= a div b) = (a <= (#0::int))"; |
|
950 |
by (stac (zdiv_zminus_zminus RS sym) 1); |
|
951 |
by (stac pos_imp_zdiv_nonneg_iff 1); |
|
952 |
by Auto_tac; |
|
953 |
qed "neg_imp_zdiv_nonneg_iff"; |
|
954 |
||
955 |
(*But not (a div b <= 0 iff a<=0); consider a=1, b=2 when a div b = 0.*) |
|
956 |
Goal "(#0::int) < b ==> (a div b < #0) = (a < #0)"; |
|
957 |
by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym, |
|
958 |
pos_imp_zdiv_nonneg_iff]) 1); |
|
959 |
qed "pos_imp_zdiv_neg_iff"; |
|
960 |
||
961 |
(*Again the law fails for <=: consider a = -1, b = -2 when a div b = 0*) |
|
962 |
Goal "b < (#0::int) ==> (a div b < #0) = (#0 < a)"; |
|
963 |
by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym, |
|
964 |
neg_imp_zdiv_nonneg_iff]) 1); |
|
965 |
qed "neg_imp_zdiv_neg_iff"; |