| author | wenzelm | 
| Fri, 23 Feb 2018 19:25:37 +0100 | |
| changeset 67710 | cc2db3239932 | 
| parent 67226 | ec32cdaab97b | 
| child 67816 | 2249b27ab1dd | 
| permissions | -rw-r--r-- | 
| 3366 | 1  | 
(* Title: HOL/Divides.thy  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
| 
6865
 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 
paulson 
parents: 
3366 
diff
changeset
 | 
3  | 
Copyright 1999 University of Cambridge  | 
| 18154 | 4  | 
*)  | 
| 3366 | 5  | 
|
| 64785 | 6  | 
section \<open>More on quotient and remainder\<close>  | 
| 3366 | 7  | 
|
| 15131 | 8  | 
theory Divides  | 
| 66817 | 9  | 
imports Parity  | 
| 15131 | 10  | 
begin  | 
| 3366 | 11  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
12  | 
subsection \<open>Numeral division with a pragmatic type class\<close>  | 
| 60758 | 13  | 
|
14  | 
text \<open>  | 
|
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
15  | 
The following type class contains everything necessary to formulate  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
16  | 
a division algorithm in ring structures with numerals, restricted  | 
| 66800 | 17  | 
to its positive segments. This is its primary motivation, and it  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
18  | 
could surely be formulated using a more fine-grained, more algebraic  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
19  | 
and less technical class hierarchy.  | 
| 60758 | 20  | 
\<close>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
21  | 
|
| 66815 | 22  | 
class unique_euclidean_semiring_numeral = semiring_parity + linordered_semidom +  | 
| 
59816
 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 
haftmann 
parents: 
59807 
diff
changeset
 | 
23  | 
assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
24  | 
and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
25  | 
and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
26  | 
and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
27  | 
and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
28  | 
and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
29  | 
and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
30  | 
and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
31  | 
assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b"  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
32  | 
fixes divmod :: "num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
33  | 
and divmod_step :: "num \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
34  | 
assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
35  | 
and divmod_step_def: "divmod_step l qr = (let (q, r) = qr  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
36  | 
in if r \<ge> numeral l then (2 * q + 1, r - numeral l)  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
37  | 
else (2 * q, r))"  | 
| 61799 | 38  | 
\<comment> \<open>These are conceptually definitions but force generated code  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
39  | 
to be monomorphic wrt. particular instances of this class which  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
40  | 
yields a significant speedup.\<close>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
41  | 
begin  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
42  | 
|
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
43  | 
lemma divmod_digit_1:  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
44  | 
assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
45  | 
shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
46  | 
and "a mod (2 * b) - b = a mod b" (is "?Q")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
47  | 
proof -  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
48  | 
from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
49  | 
by (auto intro: trans)  | 
| 60758 | 50  | 
with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive)  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
51  | 
then have [simp]: "1 \<le> a div b" by (simp add: discrete)  | 
| 60758 | 52  | 
with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)  | 
| 63040 | 53  | 
define w where "w = a div b mod 2"  | 
| 66815 | 54  | 
then have w_exhaust: "w = 0 \<or> w = 1" by auto  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
55  | 
have mod_w: "a mod (2 * b) = a mod b + b * w"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
56  | 
by (simp add: w_def mod_mult2_eq ac_simps)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
57  | 
from assms w_exhaust have "w = 1"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
58  | 
by (auto simp add: mod_w) (insert mod_less, auto)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
59  | 
with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
60  | 
have "2 * (a div (2 * b)) = a div b - w"  | 
| 64246 | 61  | 
by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)  | 
| 60758 | 62  | 
with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
63  | 
then show ?P and ?Q  | 
| 60867 | 64  | 
by (simp_all add: div mod add_implies_diff [symmetric])  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
65  | 
qed  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
66  | 
|
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
67  | 
lemma divmod_digit_0:  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
68  | 
assumes "0 < b" and "a mod (2 * b) < b"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
69  | 
shows "2 * (a div (2 * b)) = a div b" (is "?P")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
70  | 
and "a mod (2 * b) = a mod b" (is "?Q")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
71  | 
proof -  | 
| 63040 | 72  | 
define w where "w = a div b mod 2"  | 
| 66815 | 73  | 
then have w_exhaust: "w = 0 \<or> w = 1" by auto  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
74  | 
have mod_w: "a mod (2 * b) = a mod b + b * w"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
75  | 
by (simp add: w_def mod_mult2_eq ac_simps)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
76  | 
moreover have "b \<le> a mod b + b"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
77  | 
proof -  | 
| 60758 | 78  | 
from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
79  | 
then have "0 + b \<le> a mod b + b" by (rule add_right_mono)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
80  | 
then show ?thesis by simp  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
81  | 
qed  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
82  | 
moreover note assms w_exhaust  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
83  | 
ultimately have "w = 0" by auto  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
84  | 
with mod_w have mod: "a mod (2 * b) = a mod b" by simp  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
85  | 
have "2 * (a div (2 * b)) = a div b - w"  | 
| 64246 | 86  | 
by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)  | 
| 60758 | 87  | 
with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
88  | 
then show ?P and ?Q  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
89  | 
by (simp_all add: div mod)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
90  | 
qed  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
91  | 
|
| 60867 | 92  | 
lemma fst_divmod:  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
93  | 
"fst (divmod m n) = numeral m div numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
94  | 
by (simp add: divmod_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
95  | 
|
| 60867 | 96  | 
lemma snd_divmod:  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
97  | 
"snd (divmod m n) = numeral m mod numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
98  | 
by (simp add: divmod_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
99  | 
|
| 60758 | 100  | 
text \<open>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
101  | 
This is a formulation of one step (referring to one digit position)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
102  | 
in school-method division: compare the dividend at the current  | 
| 53070 | 103  | 
digit position with the remainder from previous division steps  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
104  | 
and evaluate accordingly.  | 
| 60758 | 105  | 
\<close>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
106  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
107  | 
lemma divmod_step_eq [simp]:  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
108  | 
"divmod_step l (q, r) = (if numeral l \<le> r  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
109  | 
then (2 * q + 1, r - numeral l) else (2 * q, r))"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
110  | 
by (simp add: divmod_step_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
111  | 
|
| 60758 | 112  | 
text \<open>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
113  | 
This is a formulation of school-method division.  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
114  | 
If the divisor is smaller than the dividend, terminate.  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
115  | 
If not, shift the dividend to the right until termination  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
116  | 
occurs and then reiterate single division steps in the  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
117  | 
opposite direction.  | 
| 60758 | 118  | 
\<close>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
119  | 
|
| 60867 | 120  | 
lemma divmod_divmod_step:  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
121  | 
"divmod m n = (if m < n then (0, numeral m)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
122  | 
else divmod_step n (divmod m (Num.Bit0 n)))"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
123  | 
proof (cases "m < n")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
124  | 
case True then have "numeral m < numeral n" by simp  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
125  | 
then show ?thesis  | 
| 60867 | 126  | 
by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod)  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
127  | 
next  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
128  | 
case False  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
129  | 
have "divmod m n =  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
130  | 
divmod_step n (numeral m div (2 * numeral n),  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
131  | 
numeral m mod (2 * numeral n))"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
132  | 
proof (cases "numeral n \<le> numeral m mod (2 * numeral n)")  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
133  | 
case True  | 
| 60867 | 134  | 
with divmod_step_eq  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
135  | 
have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
136  | 
(2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)"  | 
| 60867 | 137  | 
by simp  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
138  | 
moreover from True divmod_digit_1 [of "numeral m" "numeral n"]  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
139  | 
have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
140  | 
and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
141  | 
by simp_all  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
142  | 
ultimately show ?thesis by (simp only: divmod_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
143  | 
next  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
144  | 
case False then have *: "numeral m mod (2 * numeral n) < numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
145  | 
by (simp add: not_le)  | 
| 60867 | 146  | 
with divmod_step_eq  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
147  | 
have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
148  | 
(2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))"  | 
| 60867 | 149  | 
by auto  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
150  | 
moreover from * divmod_digit_0 [of "numeral n" "numeral m"]  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
151  | 
have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
152  | 
and "numeral m mod (2 * numeral n) = numeral m mod numeral n"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
153  | 
by (simp_all only: zero_less_numeral)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
154  | 
ultimately show ?thesis by (simp only: divmod_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
155  | 
qed  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
156  | 
then have "divmod m n =  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
157  | 
divmod_step n (numeral m div numeral (Num.Bit0 n),  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
158  | 
numeral m mod numeral (Num.Bit0 n))"  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
159  | 
by (simp only: numeral.simps distrib mult_1)  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
160  | 
then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))"  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
161  | 
by (simp add: divmod_def)  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
162  | 
with False show ?thesis by simp  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
163  | 
qed  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
164  | 
|
| 61799 | 165  | 
text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close>  | 
| 60867 | 166  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
167  | 
lemma divmod_trivial [simp]:  | 
| 60867 | 168  | 
"divmod Num.One Num.One = (numeral Num.One, 0)"  | 
169  | 
"divmod (Num.Bit0 m) Num.One = (numeral (Num.Bit0 m), 0)"  | 
|
170  | 
"divmod (Num.Bit1 m) Num.One = (numeral (Num.Bit1 m), 0)"  | 
|
171  | 
"divmod num.One (num.Bit0 n) = (0, Numeral1)"  | 
|
172  | 
"divmod num.One (num.Bit1 n) = (0, Numeral1)"  | 
|
173  | 
using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def)  | 
|
174  | 
||
175  | 
text \<open>Division by an even number is a right-shift\<close>  | 
|
| 58953 | 176  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
177  | 
lemma divmod_cancel [simp]:  | 
| 
53069
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
178  | 
"divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))" (is ?P)  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
179  | 
"divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))" (is ?Q)  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
180  | 
proof -  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
181  | 
have *: "\<And>q. numeral (Num.Bit0 q) = 2 * numeral q"  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
182  | 
"\<And>q. numeral (Num.Bit1 q) = 2 * numeral q + 1"  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
183  | 
by (simp_all only: numeral_mult numeral.simps distrib) simp_all  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
184  | 
have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less)  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
185  | 
then show ?P and ?Q  | 
| 60867 | 186  | 
by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1  | 
187  | 
div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2]  | 
|
188  | 
add.commute del: numeral_times_numeral)  | 
|
| 58953 | 189  | 
qed  | 
190  | 
||
| 60867 | 191  | 
text \<open>The really hard work\<close>  | 
192  | 
||
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
193  | 
lemma divmod_steps [simp]:  | 
| 60867 | 194  | 
"divmod (num.Bit0 m) (num.Bit1 n) =  | 
195  | 
(if m \<le> n then (0, numeral (num.Bit0 m))  | 
|
196  | 
else divmod_step (num.Bit1 n)  | 
|
197  | 
(divmod (num.Bit0 m)  | 
|
198  | 
(num.Bit0 (num.Bit1 n))))"  | 
|
199  | 
"divmod (num.Bit1 m) (num.Bit1 n) =  | 
|
200  | 
(if m < n then (0, numeral (num.Bit1 m))  | 
|
201  | 
else divmod_step (num.Bit1 n)  | 
|
202  | 
(divmod (num.Bit1 m)  | 
|
203  | 
(num.Bit0 (num.Bit1 n))))"  | 
|
204  | 
by (simp_all add: divmod_divmod_step)  | 
|
205  | 
||
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
206  | 
lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
207  | 
|
| 60758 | 208  | 
text \<open>Special case: divisibility\<close>  | 
| 58953 | 209  | 
|
210  | 
definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"  | 
|
211  | 
where  | 
|
212  | 
"divides_aux qr \<longleftrightarrow> snd qr = 0"  | 
|
213  | 
||
214  | 
lemma divides_aux_eq [simp]:  | 
|
215  | 
"divides_aux (q, r) \<longleftrightarrow> r = 0"  | 
|
216  | 
by (simp add: divides_aux_def)  | 
|
217  | 
||
218  | 
lemma dvd_numeral_simp [simp]:  | 
|
219  | 
"numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)"  | 
|
220  | 
by (simp add: divmod_def mod_eq_0_iff_dvd)  | 
|
| 
53069
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
221  | 
|
| 60867 | 222  | 
text \<open>Generic computation of quotient and remainder\<close>  | 
223  | 
||
224  | 
lemma numeral_div_numeral [simp]:  | 
|
225  | 
"numeral k div numeral l = fst (divmod k l)"  | 
|
226  | 
by (simp add: fst_divmod)  | 
|
227  | 
||
228  | 
lemma numeral_mod_numeral [simp]:  | 
|
229  | 
"numeral k mod numeral l = snd (divmod k l)"  | 
|
230  | 
by (simp add: snd_divmod)  | 
|
231  | 
||
232  | 
lemma one_div_numeral [simp]:  | 
|
233  | 
"1 div numeral n = fst (divmod num.One n)"  | 
|
234  | 
by (simp add: fst_divmod)  | 
|
235  | 
||
236  | 
lemma one_mod_numeral [simp]:  | 
|
237  | 
"1 mod numeral n = snd (divmod num.One n)"  | 
|
238  | 
by (simp add: snd_divmod)  | 
|
| 
64630
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
239  | 
|
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
240  | 
text \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close>  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
241  | 
|
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
242  | 
lemma cong_exp_iff_simps:  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
243  | 
"numeral n mod numeral Num.One = 0  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
244  | 
\<longleftrightarrow> True"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
245  | 
"numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
246  | 
\<longleftrightarrow> numeral n mod numeral q = 0"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
247  | 
"numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
248  | 
\<longleftrightarrow> False"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
249  | 
"numeral m mod numeral Num.One = (numeral n mod numeral Num.One)  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
250  | 
\<longleftrightarrow> True"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
251  | 
"numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
252  | 
\<longleftrightarrow> True"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
253  | 
"numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
254  | 
\<longleftrightarrow> False"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
255  | 
"numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
256  | 
\<longleftrightarrow> (numeral n mod numeral q) = 0"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
257  | 
"numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
258  | 
\<longleftrightarrow> False"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
259  | 
"numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
260  | 
\<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
261  | 
"numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
262  | 
\<longleftrightarrow> False"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
263  | 
"numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
264  | 
\<longleftrightarrow> (numeral m mod numeral q) = 0"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
265  | 
"numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
266  | 
\<longleftrightarrow> False"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
267  | 
"numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
268  | 
\<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
269  | 
by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even])  | 
| 
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
64593 
diff
changeset
 | 
270  | 
|
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
271  | 
end  | 
| 
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
272  | 
|
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
273  | 
hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
274  | 
|
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
275  | 
|
| 66817 | 276  | 
subsection \<open>More on division\<close>  | 
| 60758 | 277  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
278  | 
instantiation nat :: unique_euclidean_semiring_numeral  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
279  | 
begin  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
280  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
281  | 
definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
282  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
283  | 
divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
284  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
285  | 
definition divmod_step_nat :: "num \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
286  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
287  | 
"divmod_step_nat l qr = (let (q, r) = qr  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
288  | 
in if r \<ge> numeral l then (2 * q + 1, r - numeral l)  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
289  | 
else (2 * q, r))"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
290  | 
|
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
291  | 
instance by standard  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
292  | 
(auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq)  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
293  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
294  | 
end  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
295  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
296  | 
declare divmod_algorithm_code [where ?'a = nat, code]  | 
| 
14267
 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 
paulson 
parents: 
14208 
diff
changeset
 | 
297  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
298  | 
lemma Suc_0_div_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
299  | 
fixes k l :: num  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
300  | 
shows "Suc 0 div numeral k = fst (divmod Num.One k)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
301  | 
by (simp_all add: fst_divmod)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
302  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
303  | 
lemma Suc_0_mod_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
304  | 
fixes k l :: num  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
305  | 
shows "Suc 0 mod numeral k = snd (divmod Num.One k)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
306  | 
by (simp_all add: snd_divmod)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
307  | 
|
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
308  | 
definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
309  | 
where "divmod_nat m n = (m div n, m mod n)"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
310  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
311  | 
lemma fst_divmod_nat [simp]:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
312  | 
"fst (divmod_nat m n) = m div n"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
313  | 
by (simp add: divmod_nat_def)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
314  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
315  | 
lemma snd_divmod_nat [simp]:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
316  | 
"snd (divmod_nat m n) = m mod n"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
317  | 
by (simp add: divmod_nat_def)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
318  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
319  | 
lemma divmod_nat_if [code]:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
320  | 
"Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
321  | 
let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
322  | 
by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
323  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
324  | 
lemma [code]:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
325  | 
"m div n = fst (divmod_nat m n)"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
326  | 
"m mod n = snd (divmod_nat m n)"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
327  | 
by simp_all  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
328  | 
|
| 64635 | 329  | 
inductive eucl_rel_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool"  | 
330  | 
where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)"  | 
|
331  | 
| eucl_rel_int_dividesI: "l \<noteq> 0 \<Longrightarrow> k = q * l \<Longrightarrow> eucl_rel_int k l (q, 0)"  | 
|
332  | 
| eucl_rel_int_remainderI: "sgn r = sgn l \<Longrightarrow> \<bar>r\<bar> < \<bar>l\<bar>  | 
|
333  | 
\<Longrightarrow> k = q * l + r \<Longrightarrow> eucl_rel_int k l (q, r)"  | 
|
334  | 
||
335  | 
lemma eucl_rel_int_iff:  | 
|
336  | 
"eucl_rel_int k l (q, r) \<longleftrightarrow>  | 
|
337  | 
k = l * q + r \<and>  | 
|
338  | 
(if 0 < l then 0 \<le> r \<and> r < l else if l < 0 then l < r \<and> r \<le> 0 else q = 0)"  | 
|
339  | 
by (cases "r = 0")  | 
|
340  | 
(auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI  | 
|
341  | 
simp add: ac_simps sgn_1_pos sgn_1_neg)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
342  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
343  | 
lemma unique_quotient_lemma:  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
344  | 
"b * q' + r' \<le> b * q + r \<Longrightarrow> 0 \<le> r' \<Longrightarrow> r' < b \<Longrightarrow> r < b \<Longrightarrow> q' \<le> (q::int)"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
345  | 
apply (subgoal_tac "r' + b * (q'-q) \<le> r")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
346  | 
prefer 2 apply (simp add: right_diff_distrib)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
347  | 
apply (subgoal_tac "0 < b * (1 + q - q') ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
348  | 
apply (erule_tac [2] order_le_less_trans)  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
349  | 
prefer 2 apply (simp add: right_diff_distrib distrib_left)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
350  | 
apply (subgoal_tac "b * q' < b * (1 + q) ")  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
351  | 
prefer 2 apply (simp add: right_diff_distrib distrib_left)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
352  | 
apply (simp add: mult_less_cancel_left)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
353  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
354  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
355  | 
lemma unique_quotient_lemma_neg:  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
356  | 
"b * q' + r' \<le> b*q + r \<Longrightarrow> r \<le> 0 \<Longrightarrow> b < r \<Longrightarrow> b < r' \<Longrightarrow> q \<le> (q'::int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
357  | 
by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma) auto  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
358  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
359  | 
lemma unique_quotient:  | 
| 64635 | 360  | 
"eucl_rel_int a b (q, r) \<Longrightarrow> eucl_rel_int a b (q', r') \<Longrightarrow> q = q'"  | 
361  | 
apply (simp add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm)  | 
|
362  | 
apply (blast intro: order_antisym  | 
|
363  | 
dest: order_eq_refl [THEN unique_quotient_lemma]  | 
|
364  | 
order_eq_refl [THEN unique_quotient_lemma_neg] sym)+  | 
|
365  | 
done  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
366  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
367  | 
lemma unique_remainder:  | 
| 64635 | 368  | 
"eucl_rel_int a b (q, r) \<Longrightarrow> eucl_rel_int a b (q', r') \<Longrightarrow> r = r'"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
369  | 
apply (subgoal_tac "q = q'")  | 
| 64635 | 370  | 
apply (simp add: eucl_rel_int_iff)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
371  | 
apply (blast intro: unique_quotient)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
372  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
373  | 
|
| 64635 | 374  | 
lemma eucl_rel_int:  | 
375  | 
"eucl_rel_int k l (k div l, k mod l)"  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
376  | 
proof (cases k rule: int_cases3)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
377  | 
case zero  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
378  | 
then show ?thesis  | 
| 64635 | 379  | 
by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
380  | 
next  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
381  | 
case (pos n)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
382  | 
then show ?thesis  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
383  | 
using div_mult_mod_eq [of n]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
384  | 
by (cases l rule: int_cases3)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
385  | 
(auto simp del: of_nat_mult of_nat_add  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
386  | 
simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps  | 
| 67118 | 387  | 
eucl_rel_int_iff divide_int_def modulo_int_def)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
388  | 
next  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
389  | 
case (neg n)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
390  | 
then show ?thesis  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
391  | 
using div_mult_mod_eq [of n]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
392  | 
by (cases l rule: int_cases3)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
393  | 
(auto simp del: of_nat_mult of_nat_add  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
394  | 
simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps  | 
| 67118 | 395  | 
eucl_rel_int_iff divide_int_def modulo_int_def)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
396  | 
qed  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
397  | 
|
| 
47141
 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 
huffman 
parents: 
47140 
diff
changeset
 | 
398  | 
lemma divmod_int_unique:  | 
| 64635 | 399  | 
assumes "eucl_rel_int k l (q, r)"  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
400  | 
shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r"  | 
| 64635 | 401  | 
using assms eucl_rel_int [of k l]  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
402  | 
using unique_quotient [of k l] unique_remainder [of k l]  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
403  | 
by auto  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
404  | 
|
| 64715 | 405  | 
lemma div_abs_eq_div_nat:  | 
406  | 
"\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)"  | 
|
407  | 
by (simp add: divide_int_def)  | 
|
408  | 
||
409  | 
lemma mod_abs_eq_div_nat:  | 
|
410  | 
"\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)"  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
411  | 
by (simp add: modulo_int_def)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
412  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
413  | 
lemma zdiv_int:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
414  | 
"int (a div b) = int a div int b"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
415  | 
by (simp add: divide_int_def sgn_1_pos)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
416  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
417  | 
lemma zmod_int:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
418  | 
"int (a mod b) = int a mod int b"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
419  | 
by (simp add: modulo_int_def sgn_1_pos)  | 
| 64715 | 420  | 
|
421  | 
lemma div_sgn_abs_cancel:  | 
|
422  | 
fixes k l v :: int  | 
|
423  | 
assumes "v \<noteq> 0"  | 
|
424  | 
shows "(sgn v * \<bar>k\<bar>) div (sgn v * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"  | 
|
425  | 
proof -  | 
|
426  | 
from assms have "sgn v = - 1 \<or> sgn v = 1"  | 
|
427  | 
by (cases "v \<ge> 0") auto  | 
|
428  | 
then show ?thesis  | 
|
| 66630 | 429  | 
using assms unfolding divide_int_def [of "sgn v * \<bar>k\<bar>" "sgn v * \<bar>l\<bar>"]  | 
430  | 
by (fastforce simp add: not_less div_abs_eq_div_nat)  | 
|
| 64715 | 431  | 
qed  | 
432  | 
||
433  | 
lemma div_eq_sgn_abs:  | 
|
434  | 
fixes k l v :: int  | 
|
435  | 
assumes "sgn k = sgn l"  | 
|
436  | 
shows "k div l = \<bar>k\<bar> div \<bar>l\<bar>"  | 
|
437  | 
proof (cases "l = 0")  | 
|
438  | 
case True  | 
|
439  | 
then show ?thesis  | 
|
440  | 
by simp  | 
|
441  | 
next  | 
|
442  | 
case False  | 
|
443  | 
with assms have "(sgn k * \<bar>k\<bar>) div (sgn l * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
444  | 
using div_sgn_abs_cancel [of l k l] by simp  | 
| 64715 | 445  | 
then show ?thesis  | 
446  | 
by (simp add: sgn_mult_abs)  | 
|
447  | 
qed  | 
|
448  | 
||
449  | 
lemma div_dvd_sgn_abs:  | 
|
450  | 
fixes k l :: int  | 
|
451  | 
assumes "l dvd k"  | 
|
452  | 
shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)"  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
453  | 
proof (cases "k = 0 \<or> l = 0")  | 
| 64715 | 454  | 
case True  | 
455  | 
then show ?thesis  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
456  | 
by auto  | 
| 64715 | 457  | 
next  | 
458  | 
case False  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
459  | 
then have "k \<noteq> 0" and "l \<noteq> 0"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
460  | 
by auto  | 
| 64715 | 461  | 
show ?thesis  | 
462  | 
proof (cases "sgn l = sgn k")  | 
|
463  | 
case True  | 
|
464  | 
then show ?thesis  | 
|
465  | 
by (simp add: div_eq_sgn_abs)  | 
|
466  | 
next  | 
|
467  | 
case False  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
468  | 
with \<open>k \<noteq> 0\<close> \<open>l \<noteq> 0\<close>  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
469  | 
have "sgn l * sgn k = - 1"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
470  | 
by (simp add: sgn_if split: if_splits)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
471  | 
with assms show ?thesis  | 
| 64715 | 472  | 
unfolding divide_int_def [of k l]  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
473  | 
by (auto simp add: zdiv_int ac_simps)  | 
| 64715 | 474  | 
qed  | 
475  | 
qed  | 
|
476  | 
||
477  | 
lemma div_noneq_sgn_abs:  | 
|
478  | 
fixes k l :: int  | 
|
479  | 
assumes "l \<noteq> 0"  | 
|
480  | 
assumes "sgn k \<noteq> sgn l"  | 
|
481  | 
shows "k div l = - (\<bar>k\<bar> div \<bar>l\<bar>) - of_bool (\<not> l dvd k)"  | 
|
482  | 
using assms  | 
|
483  | 
by (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int)  | 
|
484  | 
||
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
485  | 
text\<open>Basic laws about division and remainder\<close>  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
486  | 
|
| 
47141
 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 
huffman 
parents: 
47140 
diff
changeset
 | 
487  | 
lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b"  | 
| 64635 | 488  | 
using eucl_rel_int [of a b]  | 
489  | 
by (auto simp add: eucl_rel_int_iff prod_eq_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
490  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
491  | 
lemmas pos_mod_sign = pos_mod_conj [THEN conjunct1]  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
492  | 
and pos_mod_bound = pos_mod_conj [THEN conjunct2]  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
493  | 
|
| 
47141
 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 
huffman 
parents: 
47140 
diff
changeset
 | 
494  | 
lemma neg_mod_conj: "b < (0::int) \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b"  | 
| 64635 | 495  | 
using eucl_rel_int [of a b]  | 
496  | 
by (auto simp add: eucl_rel_int_iff prod_eq_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
497  | 
|
| 45607 | 498  | 
lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1]  | 
499  | 
and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2]  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
500  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
501  | 
|
| 60758 | 502  | 
subsubsection \<open>General Properties of div and mod\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
503  | 
|
| 66886 | 504  | 
lemma div_pos_pos_trivial [simp]:  | 
505  | 
"k div l = 0" if "k \<ge> 0" and "k < l" for k l :: int  | 
|
506  | 
using that by (auto intro!: div_int_unique [of _ _ _ k] simp add: eucl_rel_int_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
507  | 
|
| 66886 | 508  | 
lemma div_neg_neg_trivial [simp]:  | 
509  | 
"k div l = 0" if "k \<le> 0" and "l < k" for k l :: int  | 
|
510  | 
using that by (auto intro!: div_int_unique [of _ _ _ k] simp add: eucl_rel_int_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
511  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
512  | 
lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1"  | 
| 
47140
 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 
huffman 
parents: 
47139 
diff
changeset
 | 
513  | 
apply (rule div_int_unique)  | 
| 64635 | 514  | 
apply (auto simp add: eucl_rel_int_iff)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
515  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
516  | 
|
| 66801 | 517  | 
lemma div_positive_int:  | 
518  | 
"k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int  | 
|
519  | 
using that by (simp add: divide_int_def div_positive)  | 
|
520  | 
||
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
521  | 
(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
522  | 
|
| 66886 | 523  | 
lemma mod_pos_pos_trivial [simp]:  | 
524  | 
"k mod l = k" if "k \<ge> 0" and "k < l" for k l :: int  | 
|
525  | 
using that by (auto intro!: mod_int_unique [of _ _ 0] simp add: eucl_rel_int_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
526  | 
|
| 66886 | 527  | 
lemma mod_neg_neg_trivial [simp]:  | 
528  | 
"k mod l = k" if "k \<le> 0" and "l < k" for k l :: int  | 
|
529  | 
using that by (auto intro!: mod_int_unique [of _ _ 0] simp add: eucl_rel_int_iff)  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
530  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
531  | 
lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b"  | 
| 
47140
 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 
huffman 
parents: 
47139 
diff
changeset
 | 
532  | 
apply (rule_tac q = "-1" in mod_int_unique)  | 
| 64635 | 533  | 
apply (auto simp add: eucl_rel_int_iff)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
534  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
535  | 
|
| 61799 | 536  | 
text\<open>There is no \<open>mod_neg_pos_trivial\<close>.\<close>  | 
| 60758 | 537  | 
|
538  | 
||
539  | 
subsubsection \<open>Laws for div and mod with Unary Minus\<close>  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
540  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
541  | 
lemma zminus1_lemma:  | 
| 64635 | 542  | 
"eucl_rel_int a b (q, r) ==> b \<noteq> 0  | 
543  | 
==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1,  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
544  | 
if r=0 then 0 else b-r)"  | 
| 66630 | 545  | 
by (force simp add: eucl_rel_int_iff right_diff_distrib)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
546  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
547  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
548  | 
lemma zdiv_zminus1_eq_if:  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
549  | 
"b \<noteq> (0::int)  | 
| 
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
550  | 
==> (-a) div b =  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
551  | 
(if a mod b = 0 then - (a div b) else - (a div b) - 1)"  | 
| 64635 | 552  | 
by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN div_int_unique])  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
553  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
554  | 
lemma zmod_zminus1_eq_if:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
555  | 
"(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
556  | 
apply (case_tac "b = 0", simp)  | 
| 64635 | 557  | 
apply (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN mod_int_unique])  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
558  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
559  | 
|
| 
64593
 
50c715579715
reoriented congruence rules in non-explosive direction
 
haftmann 
parents: 
64592 
diff
changeset
 | 
560  | 
lemma zmod_zminus1_not_zero:  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
561  | 
fixes k l :: int  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
562  | 
shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
563  | 
by (simp add: mod_eq_0_iff_dvd)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
564  | 
|
| 
64593
 
50c715579715
reoriented congruence rules in non-explosive direction
 
haftmann 
parents: 
64592 
diff
changeset
 | 
565  | 
lemma zmod_zminus2_not_zero:  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
566  | 
fixes k l :: int  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
567  | 
shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64250 
diff
changeset
 | 
568  | 
by (simp add: mod_eq_0_iff_dvd)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
569  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
570  | 
lemma zdiv_zminus2_eq_if:  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
571  | 
"b \<noteq> (0::int)  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
572  | 
==> a div (-b) =  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
573  | 
(if a mod b = 0 then - (a div b) else - (a div b) - 1)"  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
574  | 
by (auto simp add: zdiv_zminus1_eq_if div_minus_right)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
575  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
576  | 
lemma zmod_zminus2_eq_if:  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
577  | 
"a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
578  | 
by (auto simp add: zmod_zminus1_eq_if mod_minus_right)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
579  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
580  | 
|
| 60758 | 581  | 
subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
582  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
583  | 
lemma zdiv_mono1: "[| a \<le> a'; 0 < (b::int) |] ==> a div b \<le> a' div b"  | 
| 64246 | 584  | 
using mult_div_mod_eq [symmetric, of a b]  | 
585  | 
using mult_div_mod_eq [symmetric, of a' b]  | 
|
586  | 
apply -  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
587  | 
apply (rule unique_quotient_lemma)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
588  | 
apply (erule subst)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
589  | 
apply (erule subst, simp_all)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
590  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
591  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
592  | 
lemma zdiv_mono1_neg: "[| a \<le> a'; (b::int) < 0 |] ==> a' div b \<le> a div b"  | 
| 64246 | 593  | 
using mult_div_mod_eq [symmetric, of a b]  | 
594  | 
using mult_div_mod_eq [symmetric, of a' b]  | 
|
595  | 
apply -  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
596  | 
apply (rule unique_quotient_lemma_neg)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
597  | 
apply (erule subst)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
598  | 
apply (erule subst, simp_all)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
599  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
600  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
601  | 
|
| 60758 | 602  | 
subsubsection \<open>Monotonicity in the Second Argument (Divisor)\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
603  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
604  | 
lemma q_pos_lemma:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
605  | 
"[| 0 \<le> b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \<le> (q'::int)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
606  | 
apply (subgoal_tac "0 < b'* (q' + 1) ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
607  | 
apply (simp add: zero_less_mult_iff)  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
608  | 
apply (simp add: distrib_left)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
609  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
610  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
611  | 
lemma zdiv_mono2_lemma:  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
612  | 
"[| b*q + r = b'*q' + r'; 0 \<le> b'*q' + r';  | 
| 
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
613  | 
r' < b'; 0 \<le> r; 0 < b'; b' \<le> b |]  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
614  | 
==> q \<le> (q'::int)"  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
615  | 
apply (frule q_pos_lemma, assumption+)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
616  | 
apply (subgoal_tac "b*q < b* (q' + 1) ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
617  | 
apply (simp add: mult_less_cancel_left)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
618  | 
apply (subgoal_tac "b*q = r' - r + b'*q'")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
619  | 
prefer 2 apply simp  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
620  | 
apply (simp (no_asm_simp) add: distrib_left)  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
621  | 
apply (subst add.commute, rule add_less_le_mono, arith)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
622  | 
apply (rule mult_right_mono, auto)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
623  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
624  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
625  | 
lemma zdiv_mono2:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
626  | 
"[| (0::int) \<le> a; 0 < b'; b' \<le> b |] ==> a div b \<le> a div b'"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
627  | 
apply (subgoal_tac "b \<noteq> 0")  | 
| 64246 | 628  | 
prefer 2 apply arith  | 
629  | 
using mult_div_mod_eq [symmetric, of a b]  | 
|
630  | 
using mult_div_mod_eq [symmetric, of a b']  | 
|
631  | 
apply -  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
632  | 
apply (rule zdiv_mono2_lemma)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
633  | 
apply (erule subst)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
634  | 
apply (erule subst, simp_all)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
635  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
636  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
637  | 
lemma q_neg_lemma:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
638  | 
"[| b'*q' + r' < 0; 0 \<le> r'; 0 < b' |] ==> q' \<le> (0::int)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
639  | 
apply (subgoal_tac "b'*q' < 0")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
640  | 
apply (simp add: mult_less_0_iff, arith)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
641  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
642  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
643  | 
lemma zdiv_mono2_neg_lemma:  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
644  | 
"[| b*q + r = b'*q' + r'; b'*q' + r' < 0;  | 
| 
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
645  | 
r < b; 0 \<le> r'; 0 < b'; b' \<le> b |]  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
646  | 
==> q' \<le> (q::int)"  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
647  | 
apply (frule q_neg_lemma, assumption+)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
648  | 
apply (subgoal_tac "b*q' < b* (q + 1) ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
649  | 
apply (simp add: mult_less_cancel_left)  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
650  | 
apply (simp add: distrib_left)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
651  | 
apply (subgoal_tac "b*q' \<le> b'*q'")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
652  | 
prefer 2 apply (simp add: mult_right_mono_neg, arith)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
653  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
654  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
655  | 
lemma zdiv_mono2_neg:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
656  | 
"[| a < (0::int); 0 < b'; b' \<le> b |] ==> a div b' \<le> a div b"  | 
| 64246 | 657  | 
using mult_div_mod_eq [symmetric, of a b]  | 
658  | 
using mult_div_mod_eq [symmetric, of a b']  | 
|
659  | 
apply -  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
660  | 
apply (rule zdiv_mono2_neg_lemma)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
661  | 
apply (erule subst)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
662  | 
apply (erule subst, simp_all)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
663  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
664  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
665  | 
|
| 60758 | 666  | 
subsubsection \<open>More Algebraic Laws for div and mod\<close>  | 
667  | 
||
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
668  | 
lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"  | 
| 66814 | 669  | 
by (fact div_mult1_eq)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
670  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
671  | 
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
672  | 
lemma zdiv_zadd1_eq:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
673  | 
"(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"  | 
| 66814 | 674  | 
by (fact div_add1_eq)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
675  | 
|
| 67091 | 676  | 
lemma zmod_eq_0_iff: "(m mod d = 0) = (\<exists>q::int. m = d*q)"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
677  | 
by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
678  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
679  | 
(* REVISIT: should this be generalized to all semiring_div types? *)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
680  | 
lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
681  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
682  | 
|
| 60758 | 683  | 
subsubsection \<open>Proving  @{term "a div (b * c) = (a div b) div c"}\<close>
 | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
684  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
685  | 
(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
686  | 
7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
687  | 
to cause particular problems.*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
688  | 
|
| 60758 | 689  | 
text\<open>first, four lemmas to bound the remainder for the cases b<0 and b>0\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
690  | 
|
| 
55085
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
54489 
diff
changeset
 | 
691  | 
lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * c < b * (q mod c) + r"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
692  | 
apply (subgoal_tac "b * (c - q mod c) < r * 1")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
693  | 
apply (simp add: algebra_simps)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
694  | 
apply (rule order_le_less_trans)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
695  | 
apply (erule_tac [2] mult_strict_right_mono)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
696  | 
apply (rule mult_left_mono_neg)  | 
| 35216 | 697  | 
using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
698  | 
apply (simp)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
699  | 
apply (simp)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
700  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
701  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
702  | 
lemma zmult2_lemma_aux2:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
703  | 
"[| (0::int) < c; b < r; r \<le> 0 |] ==> b * (q mod c) + r \<le> 0"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
704  | 
apply (subgoal_tac "b * (q mod c) \<le> 0")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
705  | 
apply arith  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
706  | 
apply (simp add: mult_le_0_iff)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
707  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
708  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
709  | 
lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \<le> r; r < b |] ==> 0 \<le> b * (q mod c) + r"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
710  | 
apply (subgoal_tac "0 \<le> b * (q mod c) ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
711  | 
apply arith  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
712  | 
apply (simp add: zero_le_mult_iff)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
713  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
714  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
715  | 
lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
716  | 
apply (subgoal_tac "r * 1 < b * (c - q mod c) ")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
717  | 
apply (simp add: right_diff_distrib)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
718  | 
apply (rule order_less_le_trans)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
719  | 
apply (erule mult_strict_right_mono)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
720  | 
apply (rule_tac [2] mult_left_mono)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
721  | 
apply simp  | 
| 35216 | 722  | 
using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
723  | 
apply simp  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
724  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
725  | 
|
| 64635 | 726  | 
lemma zmult2_lemma: "[| eucl_rel_int a b (q, r); 0 < c |]  | 
727  | 
==> eucl_rel_int a (b * c) (q div c, b*(q mod c) + r)"  | 
|
728  | 
by (auto simp add: mult.assoc eucl_rel_int_iff linorder_neq_iff  | 
|
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
729  | 
zero_less_mult_iff distrib_left [symmetric]  | 
| 62390 | 730  | 
zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4 mult_less_0_iff split: if_split_asm)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
731  | 
|
| 53068 | 732  | 
lemma zdiv_zmult2_eq:  | 
733  | 
fixes a b c :: int  | 
|
734  | 
shows "0 \<le> c \<Longrightarrow> a div (b * c) = (a div b) div c"  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
735  | 
apply (case_tac "b = 0", simp)  | 
| 64635 | 736  | 
apply (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN div_int_unique])  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
737  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
738  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
739  | 
lemma zmod_zmult2_eq:  | 
| 53068 | 740  | 
fixes a b c :: int  | 
741  | 
shows "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
742  | 
apply (case_tac "b = 0", simp)  | 
| 64635 | 743  | 
apply (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN mod_int_unique])  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
744  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
745  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
746  | 
lemma div_pos_geq:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
747  | 
fixes k l :: int  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
748  | 
assumes "0 < l" and "l \<le> k"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
749  | 
shows "k div l = (k - l) div l + 1"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
750  | 
proof -  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
751  | 
have "k = (k - l) + l" by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
752  | 
then obtain j where k: "k = j + l" ..  | 
| 
63499
 
9c9a59949887
Tuned looping simp rules in semiring_div
 
eberlm <eberlm@in.tum.de> 
parents: 
63417 
diff
changeset
 | 
753  | 
with assms show ?thesis by (simp add: div_add_self2)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
754  | 
qed  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
755  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
756  | 
lemma mod_pos_geq:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
757  | 
fixes k l :: int  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
758  | 
assumes "0 < l" and "l \<le> k"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
759  | 
shows "k mod l = (k - l) mod l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
760  | 
proof -  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
761  | 
have "k = (k - l) + l" by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
762  | 
then obtain j where k: "k = j + l" ..  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
763  | 
with assms show ?thesis by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
764  | 
qed  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
765  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
766  | 
|
| 60758 | 767  | 
subsubsection \<open>Splitting Rules for div and mod\<close>  | 
768  | 
||
769  | 
text\<open>The proofs of the two lemmas below are essentially identical\<close>  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
770  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
771  | 
lemma split_pos_lemma:  | 
| 67091 | 772  | 
"0<k \<Longrightarrow>  | 
773  | 
P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i j)"  | 
|
| 66886 | 774  | 
by auto  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
775  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
776  | 
lemma split_neg_lemma:  | 
| 67091 | 777  | 
"k<0 \<Longrightarrow>  | 
778  | 
P(n div k :: int)(n mod k) = (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i j)"  | 
|
| 66886 | 779  | 
by auto  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
780  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
781  | 
lemma split_zdiv:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
782  | 
"P(n div k :: int) =  | 
| 67091 | 783  | 
((k = 0 \<longrightarrow> P 0) \<and>  | 
784  | 
(0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i)) \<and>  | 
|
785  | 
(k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i)))"  | 
|
| 66886 | 786  | 
apply (cases "k = 0")  | 
787  | 
apply simp  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
788  | 
apply (simp only: linorder_neq_iff)  | 
| 66886 | 789  | 
apply (auto simp add: split_pos_lemma [of concl: "%x y. P x"]  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
790  | 
split_neg_lemma [of concl: "%x y. P x"])  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
791  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
792  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
793  | 
lemma split_zmod:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
794  | 
"P(n mod k :: int) =  | 
| 67091 | 795  | 
((k = 0 \<longrightarrow> P n) \<and>  | 
796  | 
(0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P j)) \<and>  | 
|
797  | 
(k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P j)))"  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
798  | 
apply (case_tac "k=0", simp)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
799  | 
apply (simp only: linorder_neq_iff)  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
800  | 
apply (erule disjE)  | 
| 
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
801  | 
apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"]  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
802  | 
split_neg_lemma [of concl: "%x y. P y"])  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
803  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
804  | 
|
| 
63950
 
cdc1e59aa513
syntactic type class for operation mod named after mod;
 
haftmann 
parents: 
63947 
diff
changeset
 | 
805  | 
text \<open>Enable (lin)arith to deal with @{const divide} and @{const modulo}
 | 
| 
33730
 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 
webertj 
parents: 
33728 
diff
changeset
 | 
806  | 
when these are applied to some constant that is of the form  | 
| 60758 | 807  | 
  @{term "numeral k"}:\<close>
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
808  | 
declare split_zdiv [of _ _ "numeral k", arith_split] for k  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
809  | 
declare split_zmod [of _ _ "numeral k", arith_split] for k  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
810  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
811  | 
|
| 61799 | 812  | 
subsubsection \<open>Computing \<open>div\<close> and \<open>mod\<close> with shifting\<close>  | 
| 47166 | 813  | 
|
| 64635 | 814  | 
lemma pos_eucl_rel_int_mult_2:  | 
| 47166 | 815  | 
assumes "0 \<le> b"  | 
| 64635 | 816  | 
assumes "eucl_rel_int a b (q, r)"  | 
817  | 
shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)"  | 
|
818  | 
using assms unfolding eucl_rel_int_iff by auto  | 
|
819  | 
||
820  | 
lemma neg_eucl_rel_int_mult_2:  | 
|
| 47166 | 821  | 
assumes "b \<le> 0"  | 
| 64635 | 822  | 
assumes "eucl_rel_int (a + 1) b (q, r)"  | 
823  | 
shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)"  | 
|
824  | 
using assms unfolding eucl_rel_int_iff by auto  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
825  | 
|
| 60758 | 826  | 
text\<open>computing div by shifting\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
827  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
828  | 
lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"  | 
| 64635 | 829  | 
using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int]  | 
| 47166 | 830  | 
by (rule div_int_unique)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
831  | 
|
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
832  | 
lemma neg_zdiv_mult_2:  | 
| 
35815
 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 
boehmes 
parents: 
35673 
diff
changeset
 | 
833  | 
assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a"  | 
| 64635 | 834  | 
using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int]  | 
| 47166 | 835  | 
by (rule div_int_unique)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
836  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
837  | 
(* FIXME: add rules for negative numerals *)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
838  | 
lemma zdiv_numeral_Bit0 [simp]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
839  | 
"numeral (Num.Bit0 v) div numeral (Num.Bit0 w) =  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
840  | 
numeral v div (numeral w :: int)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
841  | 
unfolding numeral.simps unfolding mult_2 [symmetric]  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
842  | 
by (rule div_mult_mult1, simp)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
843  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
844  | 
lemma zdiv_numeral_Bit1 [simp]:  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
845  | 
"numeral (Num.Bit1 v) div numeral (Num.Bit0 w) =  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
846  | 
(numeral v div (numeral w :: int))"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
847  | 
unfolding numeral.simps  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
848  | 
unfolding mult_2 [symmetric] add.commute [of _ 1]  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
849  | 
by (rule pos_zdiv_mult_2, simp)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
850  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
851  | 
lemma pos_zmod_mult_2:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
852  | 
fixes a b :: int  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
853  | 
assumes "0 \<le> a"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
854  | 
shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)"  | 
| 64635 | 855  | 
using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int]  | 
| 47166 | 856  | 
by (rule mod_int_unique)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
857  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
858  | 
lemma neg_zmod_mult_2:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
859  | 
fixes a b :: int  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
860  | 
assumes "a \<le> 0"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
861  | 
shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1"  | 
| 64635 | 862  | 
using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int]  | 
| 47166 | 863  | 
by (rule mod_int_unique)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
864  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
865  | 
(* FIXME: add rules for negative numerals *)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
866  | 
lemma zmod_numeral_Bit0 [simp]:  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
867  | 
"numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) =  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
868  | 
(2::int) * (numeral v mod numeral w)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
869  | 
unfolding numeral_Bit0 [of v] numeral_Bit0 [of w]  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
870  | 
unfolding mult_2 [symmetric] by (rule mod_mult_mult1)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
871  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
872  | 
lemma zmod_numeral_Bit1 [simp]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
873  | 
"numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) =  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
874  | 
2 * (numeral v mod numeral w) + (1::int)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
875  | 
unfolding numeral_Bit1 [of v] numeral_Bit0 [of w]  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
876  | 
unfolding mult_2 [symmetric] add.commute [of _ 1]  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46560 
diff
changeset
 | 
877  | 
by (rule pos_zmod_mult_2, simp)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
878  | 
|
| 39489 | 879  | 
lemma zdiv_eq_0_iff:  | 
| 66886 | 880  | 
"i div k = 0 \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" (is "?L = ?R")  | 
881  | 
for i k :: int  | 
|
| 39489 | 882  | 
proof  | 
883  | 
assume ?L  | 
|
| 66886 | 884  | 
moreover have "?L \<longrightarrow> ?R"  | 
885  | 
by (rule split_zdiv [THEN iffD2]) simp  | 
|
886  | 
ultimately show ?R  | 
|
887  | 
by blast  | 
|
| 39489 | 888  | 
next  | 
| 66886 | 889  | 
assume ?R then show ?L  | 
890  | 
by auto  | 
|
| 39489 | 891  | 
qed  | 
892  | 
||
| 63947 | 893  | 
lemma zmod_trival_iff:  | 
894  | 
fixes i k :: int  | 
|
895  | 
shows "i mod k = i \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i"  | 
|
896  | 
proof -  | 
|
897  | 
have "i mod k = i \<longleftrightarrow> i div k = 0"  | 
|
| 
64242
 
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
 
haftmann 
parents: 
64240 
diff
changeset
 | 
898  | 
by safe (insert div_mult_mod_eq [of i k], auto)  | 
| 63947 | 899  | 
with zdiv_eq_0_iff  | 
900  | 
show ?thesis  | 
|
901  | 
by simp  | 
|
902  | 
qed  | 
|
| 39489 | 903  | 
|
| 64785 | 904  | 
|
| 60758 | 905  | 
subsubsection \<open>Quotients of Signs\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
906  | 
|
| 67083 | 907  | 
lemma div_eq_minus1: "0 < b \<Longrightarrow> - 1 div b = - 1" for b :: int  | 
908  | 
by (simp add: divide_int_def)  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
909  | 
|
| 67083 | 910  | 
lemma zmod_minus1: "0 < b \<Longrightarrow> - 1 mod b = b - 1" for b :: int  | 
911  | 
by (auto simp add: modulo_int_def)  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
912  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
913  | 
lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
914  | 
apply (subgoal_tac "a div b \<le> -1", force)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
915  | 
apply (rule order_trans)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
916  | 
apply (rule_tac a' = "-1" in zdiv_mono1)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
917  | 
apply (auto simp add: div_eq_minus1)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
918  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
919  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
920  | 
lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
921  | 
by (drule zdiv_mono1_neg, auto)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
922  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
923  | 
lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
924  | 
by (drule zdiv_mono1, auto)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
925  | 
|
| 61799 | 926  | 
text\<open>Now for some equivalences of the form \<open>a div b >=< 0 \<longleftrightarrow> \<dots>\<close>  | 
927  | 
conditional upon the sign of \<open>a\<close> or \<open>b\<close>. There are many more.  | 
|
| 60758 | 928  | 
They should all be simp rules unless that causes too much search.\<close>  | 
| 33804 | 929  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
930  | 
lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
931  | 
apply auto  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
932  | 
apply (drule_tac [2] zdiv_mono1)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
933  | 
apply (auto simp add: linorder_neq_iff)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
934  | 
apply (simp (no_asm_use) add: linorder_not_less [symmetric])  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
935  | 
apply (blast intro: div_neg_pos_less0)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
936  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
937  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
938  | 
lemma pos_imp_zdiv_pos_iff:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
939  | 
"0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
940  | 
using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k]  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
941  | 
by arith  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
942  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
943  | 
lemma neg_imp_zdiv_nonneg_iff:  | 
| 33804 | 944  | 
"b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))"  | 
| 47159 | 945  | 
apply (subst div_minus_minus [symmetric])  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
946  | 
apply (subst pos_imp_zdiv_nonneg_iff, auto)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
947  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
948  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
949  | 
(*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
950  | 
lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
951  | 
by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
952  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
953  | 
(*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
954  | 
lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
955  | 
by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
956  | 
|
| 33804 | 957  | 
lemma nonneg1_imp_zdiv_pos_iff:  | 
| 67091 | 958  | 
"(0::int) \<le> a \<Longrightarrow> (a div b > 0) = (a \<ge> b \<and> b>0)"  | 
| 33804 | 959  | 
apply rule  | 
960  | 
apply rule  | 
|
961  | 
using div_pos_pos_trivial[of a b]apply arith  | 
|
962  | 
apply(cases "b=0")apply simp  | 
|
963  | 
using div_nonneg_neg_le0[of a b]apply arith  | 
|
964  | 
using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp  | 
|
965  | 
done  | 
|
966  | 
||
| 39489 | 967  | 
lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m"  | 
968  | 
apply (rule split_zmod[THEN iffD2])  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44766 
diff
changeset
 | 
969  | 
apply(fastforce dest: q_pos_lemma intro: split_mult_pos_le)  | 
| 39489 | 970  | 
done  | 
971  | 
||
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
972  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
973  | 
subsubsection \<open>Computation of Division and Remainder\<close>  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
974  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
975  | 
instantiation int :: unique_euclidean_semiring_numeral  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
976  | 
begin  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
977  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
978  | 
definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
979  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
980  | 
"divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
981  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
982  | 
definition divmod_step_int :: "num \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
983  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
984  | 
"divmod_step_int l qr = (let (q, r) = qr  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
985  | 
in if r \<ge> numeral l then (2 * q + 1, r - numeral l)  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
986  | 
else (2 * q, r))"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
987  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
988  | 
instance  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
989  | 
by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def  | 
| 66886 | 990  | 
pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq)  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
991  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
992  | 
end  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
993  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61201 
diff
changeset
 | 
994  | 
declare divmod_algorithm_code [where ?'a = int, code]  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
995  | 
|
| 60930 | 996  | 
context  | 
997  | 
begin  | 
|
998  | 
||
999  | 
qualified definition adjust_div :: "int \<times> int \<Rightarrow> int"  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1000  | 
where  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1001  | 
"adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1002  | 
|
| 60930 | 1003  | 
qualified lemma adjust_div_eq [simp, code]:  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1004  | 
"adjust_div (q, r) = q + of_bool (r \<noteq> 0)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1005  | 
by (simp add: adjust_div_def)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1006  | 
|
| 60930 | 1007  | 
qualified definition adjust_mod :: "int \<Rightarrow> int \<Rightarrow> int"  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1008  | 
where  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1009  | 
[simp]: "adjust_mod l r = (if r = 0 then 0 else l - r)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1010  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1011  | 
lemma minus_numeral_div_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1012  | 
"- numeral m div numeral n = - (adjust_div (divmod m n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1013  | 
proof -  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1014  | 
have "int (fst (divmod m n)) = fst (divmod m n)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1015  | 
by (simp only: fst_divmod divide_int_def) auto  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1016  | 
then show ?thesis  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1017  | 
by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1018  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1019  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1020  | 
lemma minus_numeral_mod_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1021  | 
"- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)"  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1022  | 
proof (cases "snd (divmod m n) = (0::int)")  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1023  | 
case True  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1024  | 
then show ?thesis  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1025  | 
by (simp add: mod_eq_0_iff_dvd divides_aux_def)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1026  | 
next  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1027  | 
case False  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1028  | 
then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1029  | 
by (simp only: snd_divmod modulo_int_def) auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1030  | 
then show ?thesis  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1031  | 
by (simp add: divides_aux_def adjust_div_def)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1032  | 
(simp add: divides_aux_def modulo_int_def)  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1033  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1034  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1035  | 
lemma numeral_div_minus_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1036  | 
"numeral m div - numeral n = - (adjust_div (divmod m n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1037  | 
proof -  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1038  | 
have "int (fst (divmod m n)) = fst (divmod m n)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1039  | 
by (simp only: fst_divmod divide_int_def) auto  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1040  | 
then show ?thesis  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1041  | 
by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1042  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1043  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1044  | 
lemma numeral_mod_minus_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1045  | 
"numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)"  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1046  | 
proof (cases "snd (divmod m n) = (0::int)")  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1047  | 
case True  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1048  | 
then show ?thesis  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1049  | 
by (simp add: mod_eq_0_iff_dvd divides_aux_def)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1050  | 
next  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1051  | 
case False  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1052  | 
then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1053  | 
by (simp only: snd_divmod modulo_int_def) auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1054  | 
then show ?thesis  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1055  | 
by (simp add: divides_aux_def adjust_div_def)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1056  | 
(simp add: divides_aux_def modulo_int_def)  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1057  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1058  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1059  | 
lemma minus_one_div_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1060  | 
"- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1061  | 
using minus_numeral_div_numeral [of Num.One n] by simp  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1062  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1063  | 
lemma minus_one_mod_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1064  | 
"- 1 mod numeral n = adjust_mod (numeral n) (snd (divmod Num.One n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1065  | 
using minus_numeral_mod_numeral [of Num.One n] by simp  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1066  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1067  | 
lemma one_div_minus_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1068  | 
"1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1069  | 
using numeral_div_minus_numeral [of Num.One n] by simp  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1070  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1071  | 
lemma one_mod_minus_numeral [simp]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1072  | 
"1 mod - numeral n = - adjust_mod (numeral n) (snd (divmod Num.One n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1073  | 
using numeral_mod_minus_numeral [of Num.One n] by simp  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1074  | 
|
| 60930 | 1075  | 
end  | 
1076  | 
||
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1077  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1078  | 
subsubsection \<open>Further properties\<close>  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1079  | 
|
| 66817 | 1080  | 
lemma div_int_pos_iff:  | 
1081  | 
"k div l \<ge> 0 \<longleftrightarrow> k = 0 \<or> l = 0 \<or> k \<ge> 0 \<and> l \<ge> 0  | 
|
1082  | 
\<or> k < 0 \<and> l < 0"  | 
|
1083  | 
for k l :: int  | 
|
1084  | 
apply (cases "k = 0 \<or> l = 0")  | 
|
1085  | 
apply (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff)  | 
|
1086  | 
apply (rule ccontr)  | 
|
1087  | 
apply (simp add: neg_imp_zdiv_nonneg_iff)  | 
|
1088  | 
done  | 
|
1089  | 
||
1090  | 
lemma mod_int_pos_iff:  | 
|
1091  | 
"k mod l \<ge> 0 \<longleftrightarrow> l dvd k \<or> l = 0 \<and> k \<ge> 0 \<or> l > 0"  | 
|
1092  | 
for k l :: int  | 
|
1093  | 
apply (cases "l > 0")  | 
|
1094  | 
apply (simp_all add: dvd_eq_mod_eq_0)  | 
|
1095  | 
apply (use neg_mod_conj [of l k] in \<open>auto simp add: le_less not_less\<close>)  | 
|
1096  | 
done  | 
|
1097  | 
||
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1098  | 
text \<open>Simplify expresions in which div and mod combine numerical constants\<close>  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1099  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1100  | 
lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q"  | 
| 64635 | 1101  | 
by (rule div_int_unique [of a b q r]) (simp add: eucl_rel_int_iff)  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1102  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1103  | 
lemma int_div_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a div b = q"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1104  | 
by (rule div_int_unique [of a b q r],  | 
| 64635 | 1105  | 
simp add: eucl_rel_int_iff)  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1106  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1107  | 
lemma int_mod_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a mod b = r"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1108  | 
by (rule mod_int_unique [of a b q r],  | 
| 64635 | 1109  | 
simp add: eucl_rel_int_iff)  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1110  | 
|
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1111  | 
lemma int_mod_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a mod b = r"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1112  | 
by (rule mod_int_unique [of a b q r],  | 
| 64635 | 1113  | 
simp add: eucl_rel_int_iff)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1114  | 
|
| 61944 | 1115  | 
lemma abs_div: "(y::int) dvd x \<Longrightarrow> \<bar>x div y\<bar> = \<bar>x\<bar> div \<bar>y\<bar>"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1116  | 
by (unfold dvd_def, cases "y=0", auto simp add: abs_mult)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1117  | 
|
| 60758 | 1118  | 
text\<open>Suggested by Matthias Daum\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1119  | 
lemma int_power_div_base:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1120  | 
"\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1121  | 
apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1122  | 
apply (erule ssubst)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1123  | 
apply (simp only: power_add)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1124  | 
apply simp_all  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1125  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1126  | 
|
| 61799 | 1127  | 
text \<open>Distributive laws for function \<open>nat\<close>.\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1128  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1129  | 
lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1130  | 
apply (rule linorder_cases [of y 0])  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1131  | 
apply (simp add: div_nonneg_neg_le0)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1132  | 
apply simp  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1133  | 
apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1134  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1135  | 
|
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1136  | 
(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1137  | 
lemma nat_mod_distrib:  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1138  | 
"\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x mod y) = nat x mod nat y"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1139  | 
apply (case_tac "y = 0", simp)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1140  | 
apply (simp add: nat_eq_iff zmod_int)  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1141  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1142  | 
|
| 60758 | 1143  | 
text\<open>Suggested by Matthias Daum\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1144  | 
lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1145  | 
apply (subgoal_tac "nat x div nat k < nat x")  | 
| 34225 | 1146  | 
apply (simp add: nat_div_distrib [symmetric])  | 
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1147  | 
apply (rule div_less_dividend, simp_all)  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1148  | 
done  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1149  | 
|
| 66837 | 1150  | 
lemma mod_eq_dvd_iff_nat:  | 
1151  | 
"m mod q = n mod q \<longleftrightarrow> q dvd m - n" if "m \<ge> n" for m n q :: nat  | 
|
1152  | 
proof -  | 
|
1153  | 
have "int m mod int q = int n mod int q \<longleftrightarrow> int q dvd int m - int n"  | 
|
1154  | 
by (simp add: mod_eq_dvd_iff)  | 
|
1155  | 
with that have "int (m mod q) = int (n mod q) \<longleftrightarrow> int q dvd int (m - n)"  | 
|
1156  | 
by (simp only: of_nat_mod of_nat_diff)  | 
|
1157  | 
then show ?thesis  | 
|
| 67118 | 1158  | 
by simp  | 
| 66837 | 1159  | 
qed  | 
1160  | 
||
1161  | 
lemma mod_eq_nat1E:  | 
|
1162  | 
fixes m n q :: nat  | 
|
1163  | 
assumes "m mod q = n mod q" and "m \<ge> n"  | 
|
1164  | 
obtains s where "m = n + q * s"  | 
|
1165  | 
proof -  | 
|
1166  | 
from assms have "q dvd m - n"  | 
|
1167  | 
by (simp add: mod_eq_dvd_iff_nat)  | 
|
1168  | 
then obtain s where "m - n = q * s" ..  | 
|
1169  | 
with \<open>m \<ge> n\<close> have "m = n + q * s"  | 
|
1170  | 
by simp  | 
|
1171  | 
with that show thesis .  | 
|
1172  | 
qed  | 
|
1173  | 
||
1174  | 
lemma mod_eq_nat2E:  | 
|
1175  | 
fixes m n q :: nat  | 
|
1176  | 
assumes "m mod q = n mod q" and "n \<ge> m"  | 
|
1177  | 
obtains s where "n = m + q * s"  | 
|
1178  | 
using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps)  | 
|
1179  | 
||
1180  | 
lemma nat_mod_eq_lemma:  | 
|
1181  | 
assumes "(x::nat) mod n = y mod n" and "y \<le> x"  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1182  | 
shows "\<exists>q. x = y + n * q"  | 
| 66837 | 1183  | 
using assms by (rule mod_eq_nat1E) rule  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1184  | 
|
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
1185  | 
lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1186  | 
(is "?lhs = ?rhs")  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1187  | 
proof  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1188  | 
assume H: "x mod n = y mod n"  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1189  | 
  {assume xy: "x \<le> y"
 | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1190  | 
from H have th: "y mod n = x mod n" by simp  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
1191  | 
from nat_mod_eq_lemma[OF th xy] have ?rhs  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1192  | 
apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)}  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1193  | 
moreover  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1194  | 
  {assume xy: "y \<le> x"
 | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
1195  | 
from nat_mod_eq_lemma[OF H xy] have ?rhs  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1196  | 
apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)}  | 
| 
60562
 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 
paulson <lp15@cam.ac.uk> 
parents: 
60517 
diff
changeset
 | 
1197  | 
ultimately show ?rhs using linear[of x y] by blast  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1198  | 
next  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1199  | 
assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1200  | 
hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1201  | 
thus ?lhs by simp  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1202  | 
qed  | 
| 
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1203  | 
|
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1204  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1205  | 
subsubsection \<open>Dedicated simproc for calculation\<close>  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1206  | 
|
| 60758 | 1207  | 
text \<open>  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1208  | 
There is space for improvement here: the calculation itself  | 
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1209  | 
could be carried out outside the logic, and a generic simproc  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1210  | 
(simplifier setup) for generic calculation would be helpful.  | 
| 60758 | 1211  | 
\<close>  | 
| 
53067
 
ee0b7c2315d2
type class for generic division algorithm on numerals
 
haftmann 
parents: 
53066 
diff
changeset
 | 
1212  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1213  | 
simproc_setup numeral_divmod  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1214  | 
  ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
 | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1215  | 
"0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1216  | 
"0 div - 1 :: int" | "0 mod - 1 :: int" |  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1217  | 
"0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1218  | 
"0 div - numeral b :: int" | "0 mod - numeral b :: int" |  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1219  | 
"1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1220  | 
"1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1221  | 
"1 div - 1 :: int" | "1 mod - 1 :: int" |  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1222  | 
"1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1223  | 
"1 div - numeral b :: int" |"1 mod - numeral b :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1224  | 
"- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1225  | 
"- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1226  | 
"- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1227  | 
"numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1228  | 
"numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1229  | 
"numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1230  | 
"numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1231  | 
"numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1232  | 
"- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1233  | 
"- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1234  | 
"- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1235  | 
"- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" |  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1236  | 
"- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") =  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1237  | 
\<open> let  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1238  | 
    val if_cong = the (Code.get_case_cong @{theory} @{const_name If});
 | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1239  | 
fun successful_rewrite ctxt ct =  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1240  | 
let  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1241  | 
val thm = Simplifier.rewrite ctxt ct  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1242  | 
in if Thm.is_reflexive thm then NONE else SOME thm end;  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1243  | 
in fn phi =>  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1244  | 
let  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1245  | 
      val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
 | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1246  | 
one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1247  | 
one_div_minus_numeral one_mod_minus_numeral  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1248  | 
numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1249  | 
numeral_div_minus_numeral numeral_mod_minus_numeral  | 
| 60930 | 1250  | 
div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1251  | 
numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1252  | 
divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One  | 
| 60930 | 1253  | 
case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1254  | 
minus_minus numeral_times_numeral mult_zero_right mult_1_right}  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1255  | 
        @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
 | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1256  | 
fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1257  | 
(Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1258  | 
in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1259  | 
end;  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1260  | 
\<close>  | 
| 34126 | 1261  | 
|
| 35673 | 1262  | 
|
| 60758 | 1263  | 
subsubsection \<open>Code generation\<close>  | 
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1264  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1265  | 
lemma [code]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1266  | 
fixes k :: int  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1267  | 
shows  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1268  | 
"k div 0 = 0"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1269  | 
"k mod 0 = k"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1270  | 
"0 div k = 0"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1271  | 
"0 mod k = 0"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1272  | 
"k div Int.Pos Num.One = k"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1273  | 
"k mod Int.Pos Num.One = 0"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1274  | 
"k div Int.Neg Num.One = - k"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1275  | 
"k mod Int.Neg Num.One = 0"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1276  | 
"Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1277  | 
"Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)"  | 
| 60930 | 1278  | 
"Int.Neg m div Int.Pos n = - (Divides.adjust_div (divmod m n) :: int)"  | 
1279  | 
"Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)"  | 
|
1280  | 
"Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)"  | 
|
1281  | 
"Int.Pos m mod Int.Neg n = - Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)"  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1282  | 
"Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1283  | 
"Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1284  | 
by simp_all  | 
| 
53069
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
53068 
diff
changeset
 | 
1285  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52398 
diff
changeset
 | 
1286  | 
code_identifier  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52398 
diff
changeset
 | 
1287  | 
code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith  | 
| 33364 | 1288  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1289  | 
lemma dvd_eq_mod_eq_0_numeral:  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1290  | 
"numeral x dvd (numeral y :: 'a) \<longleftrightarrow> numeral y mod numeral x = (0 :: 'a::semidom_modulo)"  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1291  | 
by (fact dvd_eq_mod_eq_0)  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60867 
diff
changeset
 | 
1292  | 
|
| 64246 | 1293  | 
declare minus_div_mult_eq_mod [symmetric, nitpick_unfold]  | 
1294  | 
||
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1295  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1296  | 
subsubsection \<open>Lemmas of doubtful value\<close>  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1297  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1298  | 
lemma mod_mult_self3':  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1299  | 
"Suc (k * n + m) mod n = Suc m mod n"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1300  | 
by (fact Suc_mod_mult_self3)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1301  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1302  | 
lemma mod_Suc_eq_Suc_mod:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1303  | 
"Suc m mod n = Suc (m mod n) mod n"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1304  | 
by (simp add: mod_simps)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1305  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1306  | 
lemma div_geq:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1307  | 
"m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1308  | 
by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1309  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1310  | 
lemma mod_geq:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1311  | 
"m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1312  | 
by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1313  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1314  | 
lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1315  | 
by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1316  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1317  | 
lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1]  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1318  | 
|
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1319  | 
(*Loses information, namely we also have r<d provided d is nonzero*)  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1320  | 
lemma mod_eqD:  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1321  | 
fixes m d r q :: nat  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1322  | 
assumes "m mod d = r"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1323  | 
shows "\<exists>q. m = r + q * d"  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1324  | 
proof -  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1325  | 
from div_mult_mod_eq obtain q where "q * d + m mod d = m" by blast  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1326  | 
with assms have "m = r + q * d" by simp  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1327  | 
then show ?thesis ..  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1328  | 
qed  | 
| 
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
1329  | 
|
| 67226 | 1330  | 
lemmas even_times_iff = even_mult_iff \<comment> \<open>FIXME duplicate\<close>  | 
| 66815 | 1331  | 
|
1332  | 
lemma mod_2_not_eq_zero_eq_one_nat:  | 
|
1333  | 
fixes n :: nat  | 
|
1334  | 
shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1"  | 
|
1335  | 
by (fact not_mod_2_eq_0_eq_1)  | 
|
1336  | 
||
1337  | 
lemma even_int_iff [simp]: "even (int n) \<longleftrightarrow> even n"  | 
|
1338  | 
by (fact even_of_nat)  | 
|
1339  | 
||
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1340  | 
lemma is_unit_int:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1341  | 
"is_unit (k::int) \<longleftrightarrow> k = 1 \<or> k = - 1"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1342  | 
by auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66815 
diff
changeset
 | 
1343  | 
|
| 
33361
 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 
haftmann 
parents: 
33340 
diff
changeset
 | 
1344  | 
end  |