author | wenzelm |
Sun, 31 Dec 2023 19:24:37 +0100 | |
changeset 79409 | e1895596e1b9 |
parent 78668 | d52934f126d4 |
permissions | -rw-r--r-- |
77061
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents:
76387
diff
changeset
|
1 |
(* Title: HOL/Euclidean_Rgins.thy |
64785 | 2 |
Author: Manuel Eberl, TU Muenchen |
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
*) |
|
5 |
||
66817 | 6 |
section \<open>Division in euclidean (semi)rings\<close> |
64785 | 7 |
|
77061
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents:
76387
diff
changeset
|
8 |
theory Euclidean_Rings |
66817 | 9 |
imports Int Lattices_Big |
64785 | 10 |
begin |
11 |
||
12 |
subsection \<open>Euclidean (semi)rings with explicit division and remainder\<close> |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
13 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
14 |
class euclidean_semiring = semidom_modulo + |
64785 | 15 |
fixes euclidean_size :: "'a \<Rightarrow> nat" |
16 |
assumes size_0 [simp]: "euclidean_size 0 = 0" |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
17 |
assumes mod_size_less: |
64785 | 18 |
"b \<noteq> 0 \<Longrightarrow> euclidean_size (a mod b) < euclidean_size b" |
19 |
assumes size_mult_mono: |
|
20 |
"b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (a * b)" |
|
21 |
begin |
|
22 |
||
66840 | 23 |
lemma euclidean_size_eq_0_iff [simp]: |
24 |
"euclidean_size b = 0 \<longleftrightarrow> b = 0" |
|
25 |
proof |
|
26 |
assume "b = 0" |
|
27 |
then show "euclidean_size b = 0" |
|
28 |
by simp |
|
29 |
next |
|
30 |
assume "euclidean_size b = 0" |
|
31 |
show "b = 0" |
|
32 |
proof (rule ccontr) |
|
33 |
assume "b \<noteq> 0" |
|
34 |
with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . |
|
35 |
with \<open>euclidean_size b = 0\<close> show False |
|
36 |
by simp |
|
37 |
qed |
|
38 |
qed |
|
39 |
||
40 |
lemma euclidean_size_greater_0_iff [simp]: |
|
41 |
"euclidean_size b > 0 \<longleftrightarrow> b \<noteq> 0" |
|
42 |
using euclidean_size_eq_0_iff [symmetric, of b] by safe simp |
|
43 |
||
64785 | 44 |
lemma size_mult_mono': "b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (b * a)" |
45 |
by (subst mult.commute) (rule size_mult_mono) |
|
46 |
||
47 |
lemma dvd_euclidean_size_eq_imp_dvd: |
|
48 |
assumes "a \<noteq> 0" and "euclidean_size a = euclidean_size b" |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
49 |
and "b dvd a" |
64785 | 50 |
shows "a dvd b" |
51 |
proof (rule ccontr) |
|
52 |
assume "\<not> a dvd b" |
|
53 |
hence "b mod a \<noteq> 0" using mod_0_imp_dvd [of b a] by blast |
|
54 |
then have "b mod a \<noteq> 0" by (simp add: mod_eq_0_iff_dvd) |
|
55 |
from \<open>b dvd a\<close> have "b dvd b mod a" by (simp add: dvd_mod_iff) |
|
56 |
then obtain c where "b mod a = b * c" unfolding dvd_def by blast |
|
57 |
with \<open>b mod a \<noteq> 0\<close> have "c \<noteq> 0" by auto |
|
58 |
with \<open>b mod a = b * c\<close> have "euclidean_size (b mod a) \<ge> euclidean_size b" |
|
59 |
using size_mult_mono by force |
|
60 |
moreover from \<open>\<not> a dvd b\<close> and \<open>a \<noteq> 0\<close> |
|
61 |
have "euclidean_size (b mod a) < euclidean_size a" |
|
62 |
using mod_size_less by blast |
|
63 |
ultimately show False using \<open>euclidean_size a = euclidean_size b\<close> |
|
64 |
by simp |
|
65 |
qed |
|
66 |
||
67 |
lemma euclidean_size_times_unit: |
|
68 |
assumes "is_unit a" |
|
69 |
shows "euclidean_size (a * b) = euclidean_size b" |
|
70 |
proof (rule antisym) |
|
71 |
from assms have [simp]: "a \<noteq> 0" by auto |
|
72 |
thus "euclidean_size (a * b) \<ge> euclidean_size b" by (rule size_mult_mono') |
|
73 |
from assms have "is_unit (1 div a)" by simp |
|
74 |
hence "1 div a \<noteq> 0" by (intro notI) simp_all |
|
75 |
hence "euclidean_size (a * b) \<le> euclidean_size ((1 div a) * (a * b))" |
|
76 |
by (rule size_mult_mono') |
|
77 |
also from assms have "(1 div a) * (a * b) = b" |
|
78 |
by (simp add: algebra_simps unit_div_mult_swap) |
|
79 |
finally show "euclidean_size (a * b) \<le> euclidean_size b" . |
|
80 |
qed |
|
81 |
||
82 |
lemma euclidean_size_unit: |
|
83 |
"is_unit a \<Longrightarrow> euclidean_size a = euclidean_size 1" |
|
84 |
using euclidean_size_times_unit [of a 1] by simp |
|
85 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
86 |
lemma unit_iff_euclidean_size: |
64785 | 87 |
"is_unit a \<longleftrightarrow> euclidean_size a = euclidean_size 1 \<and> a \<noteq> 0" |
88 |
proof safe |
|
89 |
assume A: "a \<noteq> 0" and B: "euclidean_size a = euclidean_size 1" |
|
90 |
show "is_unit a" |
|
91 |
by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all |
|
92 |
qed (auto intro: euclidean_size_unit) |
|
93 |
||
94 |
lemma euclidean_size_times_nonunit: |
|
95 |
assumes "a \<noteq> 0" "b \<noteq> 0" "\<not> is_unit a" |
|
96 |
shows "euclidean_size b < euclidean_size (a * b)" |
|
97 |
proof (rule ccontr) |
|
98 |
assume "\<not>euclidean_size b < euclidean_size (a * b)" |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
99 |
with size_mult_mono'[OF assms(1), of b] |
64785 | 100 |
have eq: "euclidean_size (a * b) = euclidean_size b" by simp |
101 |
have "a * b dvd b" |
|
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74592
diff
changeset
|
102 |
by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74592
diff
changeset
|
103 |
(use assms in simp_all) |
64785 | 104 |
hence "a * b dvd 1 * b" by simp |
105 |
with \<open>b \<noteq> 0\<close> have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) |
|
106 |
with assms(3) show False by contradiction |
|
107 |
qed |
|
108 |
||
109 |
lemma dvd_imp_size_le: |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
110 |
assumes "a dvd b" "b \<noteq> 0" |
64785 | 111 |
shows "euclidean_size a \<le> euclidean_size b" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74592
diff
changeset
|
112 |
using assms by (auto simp: size_mult_mono) |
64785 | 113 |
|
114 |
lemma dvd_proper_imp_size_less: |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
115 |
assumes "a dvd b" "\<not> b dvd a" "b \<noteq> 0" |
64785 | 116 |
shows "euclidean_size a < euclidean_size b" |
117 |
proof - |
|
118 |
from assms(1) obtain c where "b = a * c" by (erule dvdE) |
|
119 |
hence z: "b = c * a" by (simp add: mult.commute) |
|
120 |
from z assms have "\<not>is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) |
|
121 |
with z assms show ?thesis |
|
122 |
by (auto intro!: euclidean_size_times_nonunit) |
|
123 |
qed |
|
124 |
||
66798 | 125 |
lemma unit_imp_mod_eq_0: |
126 |
"a mod b = 0" if "is_unit b" |
|
127 |
using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) |
|
128 |
||
69695 | 129 |
lemma mod_eq_self_iff_div_eq_0: |
130 |
"a mod b = a \<longleftrightarrow> a div b = 0" (is "?P \<longleftrightarrow> ?Q") |
|
131 |
proof |
|
132 |
assume ?P |
|
133 |
with div_mult_mod_eq [of a b] show ?Q |
|
134 |
by auto |
|
135 |
next |
|
136 |
assume ?Q |
|
137 |
with div_mult_mod_eq [of a b] show ?P |
|
138 |
by simp |
|
139 |
qed |
|
140 |
||
67051 | 141 |
lemma coprime_mod_left_iff [simp]: |
142 |
"coprime (a mod b) b \<longleftrightarrow> coprime a b" if "b \<noteq> 0" |
|
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74592
diff
changeset
|
143 |
by (rule iffI; rule coprimeI) |
67051 | 144 |
(use that in \<open>auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\<close>) |
145 |
||
146 |
lemma coprime_mod_right_iff [simp]: |
|
147 |
"coprime a (b mod a) \<longleftrightarrow> coprime a b" if "a \<noteq> 0" |
|
148 |
using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) |
|
149 |
||
64785 | 150 |
end |
151 |
||
152 |
class euclidean_ring = idom_modulo + euclidean_semiring |
|
66886 | 153 |
begin |
154 |
||
67087 | 155 |
lemma dvd_diff_commute [ac_simps]: |
66886 | 156 |
"a dvd c - b \<longleftrightarrow> a dvd b - c" |
157 |
proof - |
|
158 |
have "a dvd c - b \<longleftrightarrow> a dvd (c - b) * - 1" |
|
159 |
by (subst dvd_mult_unit_iff) simp_all |
|
160 |
then show ?thesis |
|
161 |
by simp |
|
162 |
qed |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
163 |
|
66886 | 164 |
end |
64785 | 165 |
|
66840 | 166 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
167 |
subsection \<open>Euclidean (semi)rings with cancel rules\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
168 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
169 |
class euclidean_semiring_cancel = euclidean_semiring + |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
170 |
assumes div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
171 |
and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b" |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
172 |
begin |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
173 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
174 |
lemma div_mult_self2 [simp]: |
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
175 |
assumes "b \<noteq> 0" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
176 |
shows "(a + b * c) div b = c + a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
177 |
using assms div_mult_self1 [of b a c] by (simp add: mult.commute) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
178 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
179 |
lemma div_mult_self3 [simp]: |
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
180 |
assumes "b \<noteq> 0" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
181 |
shows "(c * b + a) div b = c + a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
182 |
using assms by (simp add: add.commute) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
183 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
184 |
lemma div_mult_self4 [simp]: |
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
185 |
assumes "b \<noteq> 0" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
186 |
shows "(b * c + a) div b = c + a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
187 |
using assms by (simp add: add.commute) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
188 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
189 |
lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
190 |
proof (cases "b = 0") |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
191 |
case True then show ?thesis by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
192 |
next |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
193 |
case False |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
194 |
have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
195 |
by (simp add: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
196 |
also from False div_mult_self1 [of b a c] have |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
197 |
"\<dots> = (c + a div b) * b + (a + c * b) mod b" |
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
198 |
by (simp add: algebra_simps) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
199 |
finally have "a = a div b * b + (a + c * b) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
200 |
by (simp add: add.commute [of a] add.assoc distrib_right) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
201 |
then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
202 |
by (simp add: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
203 |
then show ?thesis by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
204 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
205 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
206 |
lemma mod_mult_self2 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
207 |
"(a + b * c) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
208 |
by (simp add: mult.commute [of b]) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
209 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
210 |
lemma mod_mult_self3 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
211 |
"(c * b + a) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
212 |
by (simp add: add.commute) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
213 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
214 |
lemma mod_mult_self4 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
215 |
"(b * c + a) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
216 |
by (simp add: add.commute) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
217 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
218 |
lemma mod_mult_self1_is_0 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
219 |
"b * a mod b = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
220 |
using mod_mult_self2 [of 0 b a] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
221 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
222 |
lemma mod_mult_self2_is_0 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
223 |
"a * b mod b = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
224 |
using mod_mult_self1 [of 0 a b] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
225 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
226 |
lemma div_add_self1: |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
227 |
assumes "b \<noteq> 0" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
228 |
shows "(b + a) div b = a div b + 1" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
229 |
using assms div_mult_self1 [of b a 1] by (simp add: add.commute) |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
230 |
|
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
231 |
lemma div_add_self2: |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
232 |
assumes "b \<noteq> 0" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
233 |
shows "(a + b) div b = a div b + 1" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
234 |
using assms div_add_self1 [of b a] by (simp add: add.commute) |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
235 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
236 |
lemma mod_add_self1 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
237 |
"(b + a) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
238 |
using mod_mult_self1 [of a 1 b] by (simp add: add.commute) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
239 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
240 |
lemma mod_add_self2 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
241 |
"(a + b) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
242 |
using mod_mult_self1 [of a 1 b] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
243 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
244 |
lemma mod_div_trivial [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
245 |
"a mod b div b = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
246 |
proof (cases "b = 0") |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
247 |
assume "b = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
248 |
thus ?thesis by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
249 |
next |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
250 |
assume "b \<noteq> 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
251 |
hence "a div b + a mod b div b = (a mod b + a div b * b) div b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
252 |
by (rule div_mult_self1 [symmetric]) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
253 |
also have "\<dots> = a div b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
254 |
by (simp only: mod_div_mult_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
255 |
also have "\<dots> = a div b + 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
256 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
257 |
finally show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
258 |
by (rule add_left_imp_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
259 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
260 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
261 |
lemma mod_mod_trivial [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
262 |
"a mod b mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
263 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
264 |
have "a mod b mod b = (a mod b + a div b * b) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
265 |
by (simp only: mod_mult_self1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
266 |
also have "\<dots> = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
267 |
by (simp only: mod_div_mult_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
268 |
finally show ?thesis . |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
269 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
270 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
271 |
lemma mod_mod_cancel: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
272 |
assumes "c dvd b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
273 |
shows "a mod b mod c = a mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
274 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
275 |
from \<open>c dvd b\<close> obtain k where "b = c * k" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
276 |
by (rule dvdE) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
277 |
have "a mod b mod c = a mod (c * k) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
278 |
by (simp only: \<open>b = c * k\<close>) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
279 |
also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
280 |
by (simp only: mod_mult_self1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
281 |
also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
282 |
by (simp only: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
283 |
also have "\<dots> = a mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
284 |
by (simp only: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
285 |
finally show ?thesis . |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
286 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
287 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
288 |
lemma div_mult_mult2 [simp]: |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
289 |
"c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
290 |
by (drule div_mult_mult1) (simp add: mult.commute) |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
291 |
|
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
292 |
lemma div_mult_mult1_if [simp]: |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
293 |
"(c * a) div (c * b) = (if c = 0 then 0 else a div b)" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
294 |
by simp_all |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
295 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
296 |
lemma mod_mult_mult1: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
297 |
"(c * a) mod (c * b) = c * (a mod b)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
298 |
proof (cases "c = 0") |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
299 |
case True then show ?thesis by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
300 |
next |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
301 |
case False |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
302 |
from div_mult_mod_eq |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
303 |
have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
304 |
with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
305 |
= c * a + c * (a mod b)" by (simp add: algebra_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
306 |
with div_mult_mod_eq show ?thesis by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
307 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
308 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
309 |
lemma mod_mult_mult2: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
310 |
"(a * c) mod (b * c) = (a mod b) * c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
311 |
using mod_mult_mult1 [of c a b] by (simp add: mult.commute) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
312 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
313 |
lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
314 |
by (fact mod_mult_mult2 [symmetric]) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
315 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
316 |
lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
317 |
by (fact mod_mult_mult1 [symmetric]) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
318 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
319 |
lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
320 |
unfolding dvd_def by (auto simp add: mod_mult_mult1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
321 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
322 |
lemma div_plus_div_distrib_dvd_left: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
323 |
"c dvd a \<Longrightarrow> (a + b) div c = a div c + b div c" |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
324 |
by (cases "c = 0") auto |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
325 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
326 |
lemma div_plus_div_distrib_dvd_right: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
327 |
"c dvd b \<Longrightarrow> (a + b) div c = a div c + b div c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
328 |
using div_plus_div_distrib_dvd_left [of c b a] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
329 |
by (simp add: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
330 |
|
71413 | 331 |
lemma sum_div_partition: |
332 |
\<open>(\<Sum>a\<in>A. f a) div b = (\<Sum>a\<in>A \<inter> {a. b dvd f a}. f a div b) + (\<Sum>a\<in>A \<inter> {a. \<not> b dvd f a}. f a) div b\<close> |
|
333 |
if \<open>finite A\<close> |
|
334 |
proof - |
|
335 |
have \<open>A = A \<inter> {a. b dvd f a} \<union> A \<inter> {a. \<not> b dvd f a}\<close> |
|
336 |
by auto |
|
337 |
then have \<open>(\<Sum>a\<in>A. f a) = (\<Sum>a\<in>A \<inter> {a. b dvd f a} \<union> A \<inter> {a. \<not> b dvd f a}. f a)\<close> |
|
338 |
by simp |
|
339 |
also have \<open>\<dots> = (\<Sum>a\<in>A \<inter> {a. b dvd f a}. f a) + (\<Sum>a\<in>A \<inter> {a. \<not> b dvd f a}. f a)\<close> |
|
340 |
using \<open>finite A\<close> by (auto intro: sum.union_inter_neutral) |
|
341 |
finally have *: \<open>sum f A = sum f (A \<inter> {a. b dvd f a}) + sum f (A \<inter> {a. \<not> b dvd f a})\<close> . |
|
342 |
define B where B: \<open>B = A \<inter> {a. b dvd f a}\<close> |
|
343 |
with \<open>finite A\<close> have \<open>finite B\<close> and \<open>a \<in> B \<Longrightarrow> b dvd f a\<close> for a |
|
344 |
by simp_all |
|
345 |
then have \<open>(\<Sum>a\<in>B. f a) div b = (\<Sum>a\<in>B. f a div b)\<close> and \<open>b dvd (\<Sum>a\<in>B. f a)\<close> |
|
346 |
by induction (simp_all add: div_plus_div_distrib_dvd_left) |
|
347 |
then show ?thesis using * |
|
348 |
by (simp add: B div_plus_div_distrib_dvd_left) |
|
349 |
qed |
|
350 |
||
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
351 |
named_theorems mod_simps |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
352 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
353 |
text \<open>Addition respects modular equivalence.\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
354 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
355 |
lemma mod_add_left_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
356 |
"(a mod c + b) mod c = (a + b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
357 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
358 |
have "(a + b) mod c = (a div c * c + a mod c + b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
359 |
by (simp only: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
360 |
also have "\<dots> = (a mod c + b + a div c * c) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
361 |
by (simp only: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
362 |
also have "\<dots> = (a mod c + b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
363 |
by (rule mod_mult_self1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
364 |
finally show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
365 |
by (rule sym) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
366 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
367 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
368 |
lemma mod_add_right_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
369 |
"(a + b mod c) mod c = (a + b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
370 |
using mod_add_left_eq [of b c a] by (simp add: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
371 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
372 |
lemma mod_add_eq: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
373 |
"(a mod c + b mod c) mod c = (a + b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
374 |
by (simp add: mod_add_left_eq mod_add_right_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
375 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
376 |
lemma mod_sum_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
377 |
"(\<Sum>i\<in>A. f i mod a) mod a = sum f A mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
378 |
proof (induct A rule: infinite_finite_induct) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
379 |
case (insert i A) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
380 |
then have "(\<Sum>i\<in>insert i A. f i mod a) mod a |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
381 |
= (f i mod a + (\<Sum>i\<in>A. f i mod a)) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
382 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
383 |
also have "\<dots> = (f i + (\<Sum>i\<in>A. f i mod a) mod a) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
384 |
by (simp add: mod_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
385 |
also have "\<dots> = (f i + (\<Sum>i\<in>A. f i) mod a) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
386 |
by (simp add: insert.hyps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
387 |
finally show ?case |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
388 |
by (simp add: insert.hyps mod_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
389 |
qed simp_all |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
390 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
391 |
lemma mod_add_cong: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
392 |
assumes "a mod c = a' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
393 |
assumes "b mod c = b' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
394 |
shows "(a + b) mod c = (a' + b') mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
395 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
396 |
have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
397 |
unfolding assms .. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
398 |
then show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
399 |
by (simp add: mod_add_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
400 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
401 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
402 |
text \<open>Multiplication respects modular equivalence.\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
403 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
404 |
lemma mod_mult_left_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
405 |
"((a mod c) * b) mod c = (a * b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
406 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
407 |
have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
408 |
by (simp only: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
409 |
also have "\<dots> = (a mod c * b + a div c * b * c) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
410 |
by (simp only: algebra_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
411 |
also have "\<dots> = (a mod c * b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
412 |
by (rule mod_mult_self1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
413 |
finally show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
414 |
by (rule sym) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
415 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
416 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
417 |
lemma mod_mult_right_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
418 |
"(a * (b mod c)) mod c = (a * b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
419 |
using mod_mult_left_eq [of b c a] by (simp add: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
420 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
421 |
lemma mod_mult_eq: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
422 |
"((a mod c) * (b mod c)) mod c = (a * b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
423 |
by (simp add: mod_mult_left_eq mod_mult_right_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
424 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
425 |
lemma mod_prod_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
426 |
"(\<Prod>i\<in>A. f i mod a) mod a = prod f A mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
427 |
proof (induct A rule: infinite_finite_induct) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
428 |
case (insert i A) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
429 |
then have "(\<Prod>i\<in>insert i A. f i mod a) mod a |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
430 |
= (f i mod a * (\<Prod>i\<in>A. f i mod a)) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
431 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
432 |
also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i mod a) mod a)) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
433 |
by (simp add: mod_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
434 |
also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i) mod a)) mod a" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
435 |
by (simp add: insert.hyps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
436 |
finally show ?case |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
437 |
by (simp add: insert.hyps mod_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
438 |
qed simp_all |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
439 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
440 |
lemma mod_mult_cong: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
441 |
assumes "a mod c = a' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
442 |
assumes "b mod c = b' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
443 |
shows "(a * b) mod c = (a' * b') mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
444 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
445 |
have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
446 |
unfolding assms .. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
447 |
then show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
448 |
by (simp add: mod_mult_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
449 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
450 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
451 |
text \<open>Exponentiation respects modular equivalence.\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
452 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
453 |
lemma power_mod [mod_simps]: |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
454 |
"((a mod b) ^ n) mod b = (a ^ n) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
455 |
proof (induct n) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
456 |
case 0 |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
457 |
then show ?case by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
458 |
next |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
459 |
case (Suc n) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
460 |
have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
461 |
by (simp add: mod_mult_right_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
462 |
with Suc show ?case |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
463 |
by (simp add: mod_mult_left_eq mod_mult_right_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
464 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
465 |
|
71413 | 466 |
lemma power_diff_power_eq: |
467 |
\<open>a ^ m div a ^ n = (if n \<le> m then a ^ (m - n) else 1 div a ^ (n - m))\<close> |
|
468 |
if \<open>a \<noteq> 0\<close> |
|
469 |
proof (cases \<open>n \<le> m\<close>) |
|
470 |
case True |
|
471 |
with that power_diff [symmetric, of a n m] show ?thesis by simp |
|
472 |
next |
|
473 |
case False |
|
474 |
then obtain q where n: \<open>n = m + Suc q\<close> |
|
475 |
by (auto simp add: not_le dest: less_imp_Suc_add) |
|
476 |
then have \<open>a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\<close> |
|
477 |
by (simp add: power_add ac_simps) |
|
478 |
moreover from that have \<open>a ^ m \<noteq> 0\<close> |
|
479 |
by simp |
|
480 |
ultimately have \<open>a ^ m div a ^ n = 1 div a ^ Suc q\<close> |
|
481 |
by (subst (asm) div_mult_mult1) simp |
|
482 |
with False n show ?thesis |
|
483 |
by simp |
|
484 |
qed |
|
485 |
||
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
486 |
end |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
487 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
488 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
489 |
class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
490 |
begin |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
491 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
492 |
subclass idom_divide .. |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
493 |
|
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
494 |
lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
495 |
using div_mult_mult1 [of "- 1" a b] by simp |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
496 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
497 |
lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
498 |
using mod_mult_mult1 [of "- 1" a b] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
499 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
500 |
lemma div_minus_right: "a div (- b) = (- a) div b" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
501 |
using div_minus_minus [of "- a" b] by simp |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
502 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
503 |
lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
504 |
using mod_minus_minus [of "- a" b] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
505 |
|
70147
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
506 |
lemma div_minus1_right [simp]: "a div (- 1) = - a" |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
507 |
using div_minus_right [of a 1] by simp |
1657688a6406
backed out a93e6472ac9c, which does not bring anything substantial: division_ring is not commutative in multiplication but semidom_divide is
haftmann
parents:
70094
diff
changeset
|
508 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
509 |
lemma mod_minus1_right [simp]: "a mod (- 1) = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
510 |
using mod_minus_right [of a 1] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
511 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
512 |
text \<open>Negation respects modular equivalence.\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
513 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
514 |
lemma mod_minus_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
515 |
"(- (a mod b)) mod b = (- a) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
516 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
517 |
have "(- a) mod b = (- (a div b * b + a mod b)) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
518 |
by (simp only: div_mult_mod_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
519 |
also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
520 |
by (simp add: ac_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
521 |
also have "\<dots> = (- (a mod b)) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
522 |
by (rule mod_mult_self1) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
523 |
finally show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
524 |
by (rule sym) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
525 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
526 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
527 |
lemma mod_minus_cong: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
528 |
assumes "a mod b = a' mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
529 |
shows "(- a) mod b = (- a') mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
530 |
proof - |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
531 |
have "(- (a mod b)) mod b = (- (a' mod b)) mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
532 |
unfolding assms .. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
533 |
then show ?thesis |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
534 |
by (simp add: mod_minus_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
535 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
536 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
537 |
text \<open>Subtraction respects modular equivalence.\<close> |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
538 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
539 |
lemma mod_diff_left_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
540 |
"(a mod c - b) mod c = (a - b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
541 |
using mod_add_cong [of a c "a mod c" "- b" "- b"] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
542 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
543 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
544 |
lemma mod_diff_right_eq [mod_simps]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
545 |
"(a - b mod c) mod c = (a - b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
546 |
using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
547 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
548 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
549 |
lemma mod_diff_eq: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
550 |
"(a mod c - b mod c) mod c = (a - b) mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
551 |
using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
552 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
553 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
554 |
lemma mod_diff_cong: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
555 |
assumes "a mod c = a' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
556 |
assumes "b mod c = b' mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
557 |
shows "(a - b) mod c = (a' - b') mod c" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
558 |
using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
559 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
560 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
561 |
lemma minus_mod_self2 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
562 |
"(a - b) mod b = a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
563 |
using mod_diff_right_eq [of a b b] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
564 |
by (simp add: mod_diff_right_eq) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
565 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
566 |
lemma minus_mod_self1 [simp]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
567 |
"(b - a) mod b = - a mod b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
568 |
using mod_add_self2 [of "- a" b] by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
569 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
570 |
lemma mod_eq_dvd_iff: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
571 |
"a mod c = b mod c \<longleftrightarrow> c dvd a - b" (is "?P \<longleftrightarrow> ?Q") |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
572 |
proof |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
573 |
assume ?P |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
574 |
then have "(a mod c - b mod c) mod c = 0" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
575 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
576 |
then show ?Q |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
577 |
by (simp add: dvd_eq_mod_eq_0 mod_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
578 |
next |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
579 |
assume ?Q |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
580 |
then obtain d where d: "a - b = c * d" .. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
581 |
then have "a = c * d + b" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
582 |
by (simp add: algebra_simps) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
583 |
then show ?P by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
584 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
585 |
|
66837 | 586 |
lemma mod_eqE: |
587 |
assumes "a mod c = b mod c" |
|
588 |
obtains d where "b = a + c * d" |
|
589 |
proof - |
|
590 |
from assms have "c dvd a - b" |
|
591 |
by (simp add: mod_eq_dvd_iff) |
|
592 |
then obtain d where "a - b = c * d" .. |
|
593 |
then have "b = a + c * - d" |
|
594 |
by (simp add: algebra_simps) |
|
595 |
with that show thesis . |
|
596 |
qed |
|
597 |
||
67051 | 598 |
lemma invertible_coprime: |
599 |
"coprime a c" if "a * b mod c = 1" |
|
600 |
by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) |
|
601 |
||
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
602 |
end |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
603 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
604 |
|
64785 | 605 |
subsection \<open>Uniquely determined division\<close> |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
606 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
607 |
class unique_euclidean_semiring = euclidean_semiring + |
76053 | 608 |
assumes euclidean_size_mult: \<open>euclidean_size (a * b) = euclidean_size a * euclidean_size b\<close> |
609 |
fixes division_segment :: \<open>'a \<Rightarrow> 'a\<close> |
|
610 |
assumes is_unit_division_segment [simp]: \<open>is_unit (division_segment a)\<close> |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
611 |
and division_segment_mult: |
76053 | 612 |
\<open>a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> division_segment (a * b) = division_segment a * division_segment b\<close> |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
613 |
and division_segment_mod: |
76053 | 614 |
\<open>b \<noteq> 0 \<Longrightarrow> \<not> b dvd a \<Longrightarrow> division_segment (a mod b) = division_segment b\<close> |
64785 | 615 |
assumes div_bounded: |
76053 | 616 |
\<open>b \<noteq> 0 \<Longrightarrow> division_segment r = division_segment b |
64785 | 617 |
\<Longrightarrow> euclidean_size r < euclidean_size b |
76053 | 618 |
\<Longrightarrow> (q * b + r) div b = q\<close> |
64785 | 619 |
begin |
620 |
||
66839 | 621 |
lemma division_segment_not_0 [simp]: |
76053 | 622 |
\<open>division_segment a \<noteq> 0\<close> |
623 |
using is_unit_division_segment [of a] is_unitE [of \<open>division_segment a\<close>] by blast |
|
624 |
||
625 |
lemma euclidean_relationI [case_names by0 divides euclidean_relation]: |
|
626 |
\<open>(a div b, a mod b) = (q, r)\<close> |
|
627 |
if by0: \<open>b = 0 \<Longrightarrow> q = 0 \<and> r = a\<close> |
|
628 |
and divides: \<open>b \<noteq> 0 \<Longrightarrow> b dvd a \<Longrightarrow> r = 0 \<and> a = q * b\<close> |
|
629 |
and euclidean_relation: \<open>b \<noteq> 0 \<Longrightarrow> \<not> b dvd a \<Longrightarrow> division_segment r = division_segment b |
|
630 |
\<and> euclidean_size r < euclidean_size b \<and> a = q * b + r\<close> |
|
631 |
proof (cases \<open>b = 0\<close>) |
|
64785 | 632 |
case True |
76053 | 633 |
with by0 show ?thesis |
634 |
by simp |
|
64785 | 635 |
next |
636 |
case False |
|
76053 | 637 |
show ?thesis |
638 |
proof (cases \<open>b dvd a\<close>) |
|
64785 | 639 |
case True |
640 |
with \<open>b \<noteq> 0\<close> divides |
|
76053 | 641 |
show ?thesis |
642 |
by simp |
|
64785 | 643 |
next |
644 |
case False |
|
76053 | 645 |
with \<open>b \<noteq> 0\<close> euclidean_relation |
646 |
have \<open>division_segment r = division_segment b\<close> |
|
647 |
\<open>euclidean_size r < euclidean_size b\<close> \<open>a = q * b + r\<close> |
|
648 |
by simp_all |
|
649 |
from \<open>b \<noteq> 0\<close> \<open>division_segment r = division_segment b\<close> |
|
650 |
\<open>euclidean_size r < euclidean_size b\<close> |
|
651 |
have \<open>(q * b + r) div b = q\<close> |
|
652 |
by (rule div_bounded) |
|
653 |
with \<open>a = q * b + r\<close> |
|
654 |
have \<open>q = a div b\<close> |
|
655 |
by simp |
|
656 |
from \<open>a = q * b + r\<close> |
|
657 |
have \<open>a div b * b + a mod b = q * b + r\<close> |
|
64785 | 658 |
by (simp add: div_mult_mod_eq) |
76053 | 659 |
with \<open>q = a div b\<close> |
660 |
have \<open>q * b + a mod b = q * b + r\<close> |
|
661 |
by simp |
|
662 |
then have \<open>r = a mod b\<close> |
|
663 |
by simp |
|
664 |
with \<open>q = a div b\<close> |
|
665 |
show ?thesis |
|
666 |
by simp |
|
64785 | 667 |
qed |
668 |
qed |
|
669 |
||
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
670 |
subclass euclidean_semiring_cancel |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
671 |
proof |
76053 | 672 |
fix a b c |
673 |
assume \<open>b \<noteq> 0\<close> |
|
674 |
have \<open>((a + c * b) div b, (a + c * b) mod b) = (c + a div b, a mod b)\<close> |
|
76245 | 675 |
proof (induction rule: euclidean_relationI) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
676 |
case by0 |
76053 | 677 |
with \<open>b \<noteq> 0\<close> |
678 |
show ?case |
|
679 |
by simp |
|
680 |
next |
|
681 |
case divides |
|
682 |
then show ?case |
|
683 |
by (simp add: algebra_simps dvd_add_left_iff) |
|
684 |
next |
|
685 |
case euclidean_relation |
|
686 |
then have \<open>\<not> b dvd a\<close> |
|
687 |
by (simp add: dvd_add_left_iff) |
|
688 |
have \<open>a mod b + (b * c + b * (a div b)) = b * c + ((a div b) * b + a mod b)\<close> |
|
689 |
by (simp add: ac_simps) |
|
690 |
with \<open>b \<noteq> 0\<close> have *: \<open>a mod b + (b * c + b * (a div b)) = b * c + a\<close> |
|
691 |
by (simp add: div_mult_mod_eq) |
|
692 |
from \<open>\<not> b dvd a\<close> euclidean_relation show ?case |
|
693 |
by (simp_all add: algebra_simps division_segment_mod mod_size_less *) |
|
694 |
qed |
|
695 |
then show \<open>(a + c * b) div b = c + a div b\<close> |
|
696 |
by simp |
|
697 |
next |
|
698 |
fix a b c |
|
699 |
assume \<open>c \<noteq> 0\<close> |
|
700 |
have \<open>((c * a) div (c * b), (c * a) mod (c * b)) = (a div b, c * (a mod b))\<close> |
|
76245 | 701 |
proof (induction rule: euclidean_relationI) |
76053 | 702 |
case by0 |
703 |
with \<open>c \<noteq> 0\<close> show ?case |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
704 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
705 |
next |
76053 | 706 |
case divides |
707 |
then show ?case |
|
708 |
by (auto simp add: algebra_simps) |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
709 |
next |
76053 | 710 |
case euclidean_relation |
711 |
then have \<open>b \<noteq> 0\<close> \<open>a mod b \<noteq> 0\<close> |
|
712 |
by (simp_all add: mod_eq_0_iff_dvd) |
|
713 |
have \<open>c * (a mod b) + b * (c * (a div b)) = c * ((a div b) * b + a mod b)\<close> |
|
714 |
by (simp add: algebra_simps) |
|
715 |
with \<open>b \<noteq> 0\<close> have *: \<open>c * (a mod b) + b * (c * (a div b)) = c * a\<close> |
|
716 |
by (simp add: div_mult_mod_eq) |
|
717 |
from \<open>b \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have \<open>euclidean_size c * euclidean_size (a mod b) |
|
718 |
< euclidean_size c * euclidean_size b\<close> |
|
719 |
using mod_size_less [of b a] by simp |
|
720 |
with euclidean_relation \<open>b \<noteq> 0\<close> \<open>a mod b \<noteq> 0\<close> show ?case |
|
721 |
by (simp add: algebra_simps division_segment_mult division_segment_mod euclidean_size_mult *) |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
722 |
qed |
76053 | 723 |
then show \<open>(c * a) div (c * b) = a div b\<close> |
724 |
by simp |
|
725 |
qed |
|
726 |
||
727 |
lemma div_eq_0_iff: |
|
728 |
\<open>a div b = 0 \<longleftrightarrow> euclidean_size a < euclidean_size b \<or> b = 0\<close> (is "_ \<longleftrightarrow> ?P") |
|
729 |
if \<open>division_segment a = division_segment b\<close> |
|
730 |
proof (cases \<open>a = 0 \<or> b = 0\<close>) |
|
731 |
case True |
|
732 |
then show ?thesis by auto |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
733 |
next |
76053 | 734 |
case False |
735 |
then have \<open>a \<noteq> 0\<close> \<open>b \<noteq> 0\<close> |
|
736 |
by simp_all |
|
737 |
have \<open>a div b = 0 \<longleftrightarrow> euclidean_size a < euclidean_size b\<close> |
|
738 |
proof |
|
739 |
assume \<open>a div b = 0\<close> |
|
740 |
then have \<open>a mod b = a\<close> |
|
741 |
using div_mult_mod_eq [of a b] by simp |
|
742 |
with \<open>b \<noteq> 0\<close> mod_size_less [of b a] |
|
743 |
show \<open>euclidean_size a < euclidean_size b\<close> |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
744 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
745 |
next |
76053 | 746 |
assume \<open>euclidean_size a < euclidean_size b\<close> |
747 |
have \<open>(a div b, a mod b) = (0, a)\<close> |
|
76245 | 748 |
proof (induction rule: euclidean_relationI) |
76053 | 749 |
case by0 |
750 |
show ?case |
|
751 |
by simp |
|
752 |
next |
|
753 |
case divides |
|
754 |
with \<open>euclidean_size a < euclidean_size b\<close> show ?case |
|
755 |
using dvd_imp_size_le [of b a] \<open>a \<noteq> 0\<close> by simp |
|
756 |
next |
|
757 |
case euclidean_relation |
|
758 |
with \<open>euclidean_size a < euclidean_size b\<close> that |
|
759 |
show ?case |
|
760 |
by simp |
|
761 |
qed |
|
762 |
then show \<open>a div b = 0\<close> |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
763 |
by simp |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
764 |
qed |
76053 | 765 |
with \<open>b \<noteq> 0\<close> show ?thesis |
766 |
by simp |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
767 |
qed |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
768 |
|
66814 | 769 |
lemma div_mult1_eq: |
76053 | 770 |
\<open>(a * b) div c = a * (b div c) + a * (b mod c) div c\<close> |
771 |
proof - |
|
772 |
have *: \<open>(a * b) mod c + (a * (c * (b div c)) + c * (a * (b mod c) div c)) = a * b\<close> (is \<open>?A + (?B + ?C) = _\<close>) |
|
773 |
proof - |
|
774 |
have \<open>?A = a * (b mod c) mod c\<close> |
|
775 |
by (simp add: mod_mult_right_eq) |
|
776 |
then have \<open>?C + ?A = a * (b mod c)\<close> |
|
777 |
by (simp add: mult_div_mod_eq) |
|
778 |
then have \<open>?B + (?C + ?A) = a * (c * (b div c) + (b mod c))\<close> |
|
779 |
by (simp add: algebra_simps) |
|
780 |
also have \<open>\<dots> = a * b\<close> |
|
781 |
by (simp add: mult_div_mod_eq) |
|
782 |
finally show ?thesis |
|
783 |
by (simp add: algebra_simps) |
|
66814 | 784 |
qed |
76053 | 785 |
have \<open>((a * b) div c, (a * b) mod c) = (a * (b div c) + a * (b mod c) div c, (a * b) mod c)\<close> |
76245 | 786 |
proof (induction rule: euclidean_relationI) |
76053 | 787 |
case by0 |
788 |
then show ?case by simp |
|
789 |
next |
|
790 |
case divides |
|
791 |
with * show ?case |
|
792 |
by (simp add: algebra_simps) |
|
793 |
next |
|
794 |
case euclidean_relation |
|
795 |
with * show ?case |
|
796 |
by (simp add: division_segment_mod mod_size_less algebra_simps) |
|
797 |
qed |
|
66814 | 798 |
then show ?thesis |
799 |
by simp |
|
800 |
qed |
|
801 |
||
802 |
lemma div_add1_eq: |
|
76053 | 803 |
\<open>(a + b) div c = a div c + b div c + (a mod c + b mod c) div c\<close> |
804 |
proof - |
|
805 |
have *: \<open>(a + b) mod c + (c * (a div c) + (c * (b div c) + c * ((a mod c + b mod c) div c))) = a + b\<close> |
|
806 |
(is \<open>?A + (?B + (?C + ?D)) = _\<close>) |
|
807 |
proof - |
|
808 |
have \<open>?A + (?B + (?C + ?D)) = ?A + ?D + (?B + ?C)\<close> |
|
809 |
by (simp add: ac_simps) |
|
810 |
also have \<open>?A + ?D = (a mod c + b mod c) mod c + ?D\<close> |
|
811 |
by (simp add: mod_add_eq) |
|
812 |
also have \<open>\<dots> = a mod c + b mod c\<close> |
|
813 |
by (simp add: mod_mult_div_eq) |
|
814 |
finally have \<open>?A + (?B + (?C + ?D)) = (a mod c + ?B) + (b mod c + ?C)\<close> |
|
815 |
by (simp add: ac_simps) |
|
816 |
then show ?thesis |
|
817 |
by (simp add: mod_mult_div_eq) |
|
818 |
qed |
|
819 |
have \<open>((a + b) div c, (a + b) mod c) = (a div c + b div c + (a mod c + b mod c) div c, (a + b) mod c)\<close> |
|
76245 | 820 |
proof (induction rule: euclidean_relationI) |
76053 | 821 |
case by0 |
822 |
then show ?case |
|
823 |
by simp |
|
824 |
next |
|
825 |
case divides |
|
826 |
with * show ?case |
|
66814 | 827 |
by (simp add: algebra_simps) |
76053 | 828 |
next |
829 |
case euclidean_relation |
|
830 |
with * show ?case |
|
831 |
by (simp add: division_segment_mod mod_size_less algebra_simps) |
|
66814 | 832 |
qed |
833 |
then show ?thesis |
|
834 |
by simp |
|
835 |
qed |
|
836 |
||
64785 | 837 |
end |
838 |
||
839 |
class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
840 |
begin |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
841 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
842 |
subclass euclidean_ring_cancel .. |
64785 | 843 |
|
844 |
end |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
845 |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
846 |
|
76387 | 847 |
subsection \<open>Division on \<^typ>\<open>nat\<close>\<close> |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
848 |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
849 |
instantiation nat :: normalization_semidom |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
850 |
begin |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
851 |
|
76053 | 852 |
definition normalize_nat :: \<open>nat \<Rightarrow> nat\<close> |
853 |
where [simp]: \<open>normalize = (id :: nat \<Rightarrow> nat)\<close> |
|
854 |
||
855 |
definition unit_factor_nat :: \<open>nat \<Rightarrow> nat\<close> |
|
856 |
where \<open>unit_factor n = of_bool (n > 0)\<close> for n :: nat |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
857 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
858 |
lemma unit_factor_simps [simp]: |
76053 | 859 |
\<open>unit_factor 0 = (0::nat)\<close> |
860 |
\<open>unit_factor (Suc n) = 1\<close> |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
861 |
by (simp_all add: unit_factor_nat_def) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
862 |
|
76053 | 863 |
definition divide_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> |
864 |
where \<open>m div n = (if n = 0 then 0 else Max {k. k * n \<le> m})\<close> for m n :: nat |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
865 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
866 |
instance |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
867 |
by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
868 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
869 |
end |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
870 |
|
67051 | 871 |
lemma coprime_Suc_0_left [simp]: |
872 |
"coprime (Suc 0) n" |
|
873 |
using coprime_1_left [of n] by simp |
|
874 |
||
875 |
lemma coprime_Suc_0_right [simp]: |
|
876 |
"coprime n (Suc 0)" |
|
877 |
using coprime_1_right [of n] by simp |
|
878 |
||
879 |
lemma coprime_common_divisor_nat: "coprime a b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> x = 1" |
|
880 |
for a b :: nat |
|
881 |
by (drule coprime_common_divisor [of _ _ x]) simp_all |
|
882 |
||
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
883 |
instantiation nat :: unique_euclidean_semiring |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
884 |
begin |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
885 |
|
76053 | 886 |
definition euclidean_size_nat :: \<open>nat \<Rightarrow> nat\<close> |
887 |
where [simp]: \<open>euclidean_size_nat = id\<close> |
|
888 |
||
889 |
definition division_segment_nat :: \<open>nat \<Rightarrow> nat\<close> |
|
890 |
where [simp]: \<open>division_segment n = 1\<close> for n :: nat |
|
891 |
||
892 |
definition modulo_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> |
|
893 |
where \<open>m mod n = m - (m div n * n)\<close> for m n :: nat |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
894 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
895 |
instance proof |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
896 |
fix m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
897 |
have ex: "\<exists>k. k * n \<le> l" for l :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
898 |
by (rule exI [of _ 0]) simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
899 |
have fin: "finite {k. k * n \<le> l}" if "n > 0" for l |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
900 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
901 |
from that have "{k. k * n \<le> l} \<subseteq> {k. k \<le> l}" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
902 |
by (cases n) auto |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
903 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
904 |
by (rule finite_subset) simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
905 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
906 |
have mult_div_unfold: "n * (m div n) = Max {l. l \<le> m \<and> n dvd l}" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
907 |
proof (cases "n = 0") |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
908 |
case True |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
909 |
moreover have "{l. l = 0 \<and> l \<le> m} = {0::nat}" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
910 |
by auto |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
911 |
ultimately show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
912 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
913 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
914 |
case False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
915 |
with ex [of m] fin have "n * Max {k. k * n \<le> m} = Max (times n ` {k. k * n \<le> m})" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
916 |
by (auto simp add: nat_mult_max_right intro: hom_Max_commute) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
917 |
also have "times n ` {k. k * n \<le> m} = {l. l \<le> m \<and> n dvd l}" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
918 |
by (auto simp add: ac_simps elim!: dvdE) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
919 |
finally show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
920 |
using False by (simp add: divide_nat_def ac_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
921 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
922 |
have less_eq: "m div n * n \<le> m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
923 |
by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
924 |
then show "m div n * n + m mod n = m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
925 |
by (simp add: modulo_nat_def) |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
926 |
assume "n \<noteq> 0" |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
927 |
show "euclidean_size (m mod n) < euclidean_size n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
928 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
929 |
have "m < Suc (m div n) * n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
930 |
proof (rule ccontr) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
931 |
assume "\<not> m < Suc (m div n) * n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
932 |
then have "Suc (m div n) * n \<le> m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
933 |
by (simp add: not_less) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
934 |
moreover from \<open>n \<noteq> 0\<close> have "Max {k. k * n \<le> m} < Suc (m div n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
935 |
by (simp add: divide_nat_def) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
936 |
with \<open>n \<noteq> 0\<close> ex fin have "\<And>k. k * n \<le> m \<Longrightarrow> k < Suc (m div n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
937 |
by auto |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
938 |
ultimately have "Suc (m div n) < Suc (m div n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
939 |
by blast |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
940 |
then show False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
941 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
942 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
943 |
with \<open>n \<noteq> 0\<close> show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
944 |
by (simp add: modulo_nat_def) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
945 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
946 |
show "euclidean_size m \<le> euclidean_size (m * n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
947 |
using \<open>n \<noteq> 0\<close> by (cases n) simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
948 |
fix q r :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
949 |
show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
950 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
951 |
from that have "r < n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
952 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
953 |
have "k \<le> q" if "k * n \<le> q * n + r" for k |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
954 |
proof (rule ccontr) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
955 |
assume "\<not> k \<le> q" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
956 |
then have "q < k" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
957 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
958 |
then obtain l where "k = Suc (q + l)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
959 |
by (auto simp add: less_iff_Suc_add) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
960 |
with \<open>r < n\<close> that show False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
961 |
by (simp add: algebra_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
962 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
963 |
with \<open>n \<noteq> 0\<close> ex fin show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
964 |
by (auto simp add: divide_nat_def Max_eq_iff) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
965 |
qed |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
966 |
qed simp_all |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
967 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66798
diff
changeset
|
968 |
end |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
969 |
|
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
970 |
lemma euclidean_relation_natI [case_names by0 divides euclidean_relation]: |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
971 |
\<open>(m div n, m mod n) = (q, r)\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
972 |
if by0: \<open>n = 0 \<Longrightarrow> q = 0 \<and> r = m\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
973 |
and divides: \<open>n > 0 \<Longrightarrow> n dvd m \<Longrightarrow> r = 0 \<and> m = q * n\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
974 |
and euclidean_relation: \<open>n > 0 \<Longrightarrow> \<not> n dvd m \<Longrightarrow> r < n \<and> m = q * n + r\<close> for m n q r :: nat |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
975 |
by (rule euclidean_relationI) (use that in simp_all) |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
976 |
|
75881 | 977 |
lemma div_nat_eqI: |
76053 | 978 |
\<open>m div n = q\<close> if \<open>n * q \<le> m\<close> and \<open>m < n * Suc q\<close> for m n q :: nat |
979 |
proof - |
|
980 |
have \<open>(m div n, m mod n) = (q, m - n * q)\<close> |
|
76245 | 981 |
proof (induction rule: euclidean_relation_natI) |
76053 | 982 |
case by0 |
983 |
with that show ?case |
|
984 |
by simp |
|
985 |
next |
|
986 |
case divides |
|
987 |
from \<open>n dvd m\<close> obtain s where \<open>m = n * s\<close> .. |
|
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
988 |
with \<open>n > 0\<close> that have \<open>s < Suc q\<close> |
76053 | 989 |
by (simp only: mult_less_cancel1) |
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
990 |
with \<open>m = n * s\<close> \<open>n > 0\<close> that have \<open>q = s\<close> |
76053 | 991 |
by simp |
992 |
with \<open>m = n * s\<close> show ?case |
|
993 |
by (simp add: ac_simps) |
|
994 |
next |
|
995 |
case euclidean_relation |
|
996 |
with that show ?case |
|
997 |
by (simp add: ac_simps) |
|
998 |
qed |
|
999 |
then show ?thesis |
|
1000 |
by simp |
|
1001 |
qed |
|
75881 | 1002 |
|
1003 |
lemma mod_nat_eqI: |
|
76053 | 1004 |
\<open>m mod n = r\<close> if \<open>r < n\<close> and \<open>r \<le> m\<close> and \<open>n dvd m - r\<close> for m n r :: nat |
1005 |
proof - |
|
1006 |
have \<open>(m div n, m mod n) = ((m - r) div n, r)\<close> |
|
76245 | 1007 |
proof (induction rule: euclidean_relation_natI) |
76053 | 1008 |
case by0 |
1009 |
with that show ?case |
|
1010 |
by simp |
|
1011 |
next |
|
1012 |
case divides |
|
1013 |
from that dvd_minus_add [of r \<open>m\<close> 1 n] |
|
1014 |
have \<open>n dvd m + (n - r)\<close> |
|
1015 |
by simp |
|
1016 |
with divides have \<open>n dvd n - r\<close> |
|
1017 |
by (simp add: dvd_add_right_iff) |
|
1018 |
then have \<open>n \<le> n - r\<close> |
|
1019 |
by (rule dvd_imp_le) (use \<open>r < n\<close> in simp) |
|
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1020 |
with \<open>n > 0\<close> have \<open>r = 0\<close> |
76053 | 1021 |
by simp |
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1022 |
with \<open>n > 0\<close> that show ?case |
76053 | 1023 |
by simp |
1024 |
next |
|
1025 |
case euclidean_relation |
|
1026 |
with that show ?case |
|
1027 |
by (simp add: ac_simps) |
|
1028 |
qed |
|
1029 |
then show ?thesis |
|
1030 |
by simp |
|
1031 |
qed |
|
75881 | 1032 |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1033 |
text \<open>Tool support\<close> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1034 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1035 |
ML \<open> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1036 |
structure Cancel_Div_Mod_Nat = Cancel_Div_Mod |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1037 |
( |
69593 | 1038 |
val div_name = \<^const_name>\<open>divide\<close>; |
1039 |
val mod_name = \<^const_name>\<open>modulo\<close>; |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1040 |
val mk_binop = HOLogic.mk_binop; |
69593 | 1041 |
val dest_plus = HOLogic.dest_bin \<^const_name>\<open>Groups.plus\<close> HOLogic.natT; |
66813 | 1042 |
val mk_sum = Arith_Data.mk_sum; |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1043 |
fun dest_sum tm = |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1044 |
if HOLogic.is_zero tm then [] |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1045 |
else |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1046 |
(case try HOLogic.dest_Suc tm of |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1047 |
SOME t => HOLogic.Suc_zero :: dest_sum t |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1048 |
| NONE => |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1049 |
(case try dest_plus tm of |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1050 |
SOME (t, u) => dest_sum t @ dest_sum u |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1051 |
| NONE => [tm])); |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1052 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1053 |
val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1054 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1055 |
val prove_eq_sums = Arith_Data.prove_conv2 all_tac |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1056 |
(Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1057 |
) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1058 |
\<close> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1059 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1060 |
simproc_setup cancel_div_mod_nat ("(m::nat) + n") = |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1061 |
\<open>K Cancel_Div_Mod_Nat.proc\<close> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1062 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1063 |
lemma div_mult_self_is_m [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1064 |
"m * n div n = m" if "n > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1065 |
using that by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1066 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1067 |
lemma div_mult_self1_is_m [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1068 |
"n * m div n = m" if "n > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1069 |
using that by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1070 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1071 |
lemma mod_less_divisor [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1072 |
"m mod n < n" if "n > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1073 |
using mod_size_less [of n m] that by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1074 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1075 |
lemma mod_le_divisor [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1076 |
"m mod n \<le> n" if "n > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1077 |
using that by (auto simp add: le_less) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1078 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1079 |
lemma div_times_less_eq_dividend [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1080 |
"m div n * n \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1081 |
by (simp add: minus_mod_eq_div_mult [symmetric]) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1082 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1083 |
lemma times_div_less_eq_dividend [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1084 |
"n * (m div n) \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1085 |
using div_times_less_eq_dividend [of m n] |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1086 |
by (simp add: ac_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1087 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1088 |
lemma dividend_less_div_times: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1089 |
"m < n + (m div n) * n" if "0 < n" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1090 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1091 |
from that have "m mod n < n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1092 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1093 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1094 |
by (simp add: minus_mod_eq_div_mult [symmetric]) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1095 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1096 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1097 |
lemma dividend_less_times_div: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1098 |
"m < n + n * (m div n)" if "0 < n" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1099 |
using dividend_less_div_times [of n m] that |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1100 |
by (simp add: ac_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1101 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1102 |
lemma mod_Suc_le_divisor [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1103 |
"m mod Suc n \<le> n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1104 |
using mod_less_divisor [of "Suc n" m] by arith |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1105 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1106 |
lemma mod_less_eq_dividend [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1107 |
"m mod n \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1108 |
proof (rule add_leD2) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1109 |
from div_mult_mod_eq have "m div n * n + m mod n = m" . |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1110 |
then show "m div n * n + m mod n \<le> m" by auto |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1111 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1112 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1113 |
lemma |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1114 |
div_less [simp]: "m div n = 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1115 |
and mod_less [simp]: "m mod n = m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1116 |
if "m < n" for m n :: nat |
76053 | 1117 |
using that by (auto intro: div_nat_eqI mod_nat_eqI) |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1118 |
|
75881 | 1119 |
lemma split_div: |
1120 |
\<open>P (m div n) \<longleftrightarrow> |
|
1121 |
(n = 0 \<longrightarrow> P 0) \<and> |
|
1122 |
(n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> P i))\<close> (is ?div) |
|
1123 |
and split_mod: |
|
1124 |
\<open>Q (m mod n) \<longleftrightarrow> |
|
1125 |
(n = 0 \<longrightarrow> Q m) \<and> |
|
1126 |
(n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> Q j))\<close> (is ?mod) |
|
1127 |
for m n :: nat |
|
1128 |
proof - |
|
1129 |
have *: \<open>R (m div n) (m mod n) \<longleftrightarrow> |
|
1130 |
(n = 0 \<longrightarrow> R 0 m) \<and> |
|
1131 |
(n \<noteq> 0 \<longrightarrow> (\<forall>i j. j < n \<and> m = n * i + j \<longrightarrow> R i j))\<close> for R |
|
1132 |
by (cases \<open>n = 0\<close>) auto |
|
1133 |
from * [of \<open>\<lambda>q _. P q\<close>] show ?div . |
|
1134 |
from * [of \<open>\<lambda>_ r. Q r\<close>] show ?mod . |
|
1135 |
qed |
|
1136 |
||
1137 |
declare split_div [of _ _ \<open>numeral n\<close>, linarith_split] for n |
|
1138 |
declare split_mod [of _ _ \<open>numeral n\<close>, linarith_split] for n |
|
1139 |
||
1140 |
lemma split_div': |
|
1141 |
"P (m div n) \<longleftrightarrow> n = 0 \<and> P 0 \<or> (\<exists>q. (n * q \<le> m \<and> m < n * Suc q) \<and> P q)" |
|
1142 |
proof (cases "n = 0") |
|
1143 |
case True |
|
1144 |
then show ?thesis |
|
1145 |
by simp |
|
1146 |
next |
|
1147 |
case False |
|
1148 |
then have "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> m div n = q" for q |
|
1149 |
by (auto intro: div_nat_eqI dividend_less_times_div) |
|
1150 |
then show ?thesis |
|
1151 |
by auto |
|
1152 |
qed |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1153 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1154 |
lemma le_div_geq: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1155 |
"m div n = Suc ((m - n) div n)" if "0 < n" and "n \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1156 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1157 |
from \<open>n \<le> m\<close> obtain q where "m = n + q" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1158 |
by (auto simp add: le_iff_add) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1159 |
with \<open>0 < n\<close> show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1160 |
by (simp add: div_add_self1) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1161 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1162 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1163 |
lemma le_mod_geq: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1164 |
"m mod n = (m - n) mod n" if "n \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1165 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1166 |
from \<open>n \<le> m\<close> obtain q where "m = n + q" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1167 |
by (auto simp add: le_iff_add) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1168 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1169 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1170 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1171 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1172 |
lemma div_if: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1173 |
"m div n = (if m < n \<or> n = 0 then 0 else Suc ((m - n) div n))" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1174 |
by (simp add: le_div_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1175 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1176 |
lemma mod_if: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1177 |
"m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1178 |
by (simp add: le_mod_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1179 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1180 |
lemma div_eq_0_iff: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1181 |
"m div n = 0 \<longleftrightarrow> m < n \<or> n = 0" for m n :: nat |
66886 | 1182 |
by (simp add: div_eq_0_iff) |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1183 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1184 |
lemma div_greater_zero_iff: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1185 |
"m div n > 0 \<longleftrightarrow> n \<le> m \<and> n > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1186 |
using div_eq_0_iff [of m n] by auto |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1187 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1188 |
lemma mod_greater_zero_iff_not_dvd: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1189 |
"m mod n > 0 \<longleftrightarrow> \<not> n dvd m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1190 |
by (simp add: dvd_eq_mod_eq_0) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1191 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1192 |
lemma div_by_Suc_0 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1193 |
"m div Suc 0 = m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1194 |
using div_by_1 [of m] by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1195 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1196 |
lemma mod_by_Suc_0 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1197 |
"m mod Suc 0 = 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1198 |
using mod_by_1 [of m] by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1199 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1200 |
lemma div2_Suc_Suc [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1201 |
"Suc (Suc m) div 2 = Suc (m div 2)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1202 |
by (simp add: numeral_2_eq_2 le_div_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1203 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1204 |
lemma Suc_n_div_2_gt_zero [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1205 |
"0 < Suc n div 2" if "n > 0" for n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1206 |
using that by (cases n) simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1207 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1208 |
lemma div_2_gt_zero [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1209 |
"0 < n div 2" if "Suc 0 < n" for n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1210 |
using that Suc_n_div_2_gt_zero [of "n - 1"] by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1211 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1212 |
lemma mod2_Suc_Suc [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1213 |
"Suc (Suc m) mod 2 = m mod 2" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1214 |
by (simp add: numeral_2_eq_2 le_mod_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1215 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1216 |
lemma add_self_div_2 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1217 |
"(m + m) div 2 = m" for m :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1218 |
by (simp add: mult_2 [symmetric]) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1219 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1220 |
lemma add_self_mod_2 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1221 |
"(m + m) mod 2 = 0" for m :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1222 |
by (simp add: mult_2 [symmetric]) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1223 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1224 |
lemma mod2_gr_0 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1225 |
"0 < m mod 2 \<longleftrightarrow> m mod 2 = 1" for m :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1226 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1227 |
have "m mod 2 < 2" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1228 |
by (rule mod_less_divisor) simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1229 |
then have "m mod 2 = 0 \<or> m mod 2 = 1" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1230 |
by arith |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1231 |
then show ?thesis |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1232 |
by auto |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1233 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1234 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1235 |
lemma mod_Suc_eq [mod_simps]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1236 |
"Suc (m mod n) mod n = Suc m mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1237 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1238 |
have "(m mod n + 1) mod n = (m + 1) mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1239 |
by (simp only: mod_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1240 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1241 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1242 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1243 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1244 |
lemma mod_Suc_Suc_eq [mod_simps]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1245 |
"Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1246 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1247 |
have "(m mod n + 2) mod n = (m + 2) mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1248 |
by (simp only: mod_simps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1249 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1250 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1251 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1252 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1253 |
lemma |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1254 |
Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1255 |
and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1256 |
and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1257 |
and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1258 |
by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1259 |
|
67083 | 1260 |
lemma Suc_0_mod_eq [simp]: |
1261 |
"Suc 0 mod n = of_bool (n \<noteq> Suc 0)" |
|
1262 |
by (cases n) simp_all |
|
1263 |
||
76053 | 1264 |
lemma div_mult2_eq: |
1265 |
\<open>m div (n * q) = (m div n) div q\<close> (is ?Q) |
|
1266 |
and mod_mult2_eq: |
|
1267 |
\<open>m mod (n * q) = n * (m div n mod q) + m mod n\<close> (is ?R) |
|
1268 |
for m n q :: nat |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1269 |
proof - |
76053 | 1270 |
have \<open>(m div (n * q), m mod (n * q)) = ((m div n) div q, n * (m div n mod q) + m mod n)\<close> |
76245 | 1271 |
proof (induction rule: euclidean_relation_natI) |
76053 | 1272 |
case by0 |
1273 |
then show ?case |
|
1274 |
by auto |
|
1275 |
next |
|
1276 |
case divides |
|
1277 |
from \<open>n * q dvd m\<close> obtain t where \<open>m = n * q * t\<close> .. |
|
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1278 |
with \<open>n * q > 0\<close> show ?case |
76053 | 1279 |
by (simp add: algebra_simps) |
1280 |
next |
|
1281 |
case euclidean_relation |
|
1282 |
then have \<open>n > 0\<close> \<open>q > 0\<close> |
|
1283 |
by simp_all |
|
1284 |
from \<open>n > 0\<close> have \<open>m mod n < n\<close> |
|
1285 |
by (rule mod_less_divisor) |
|
1286 |
from \<open>q > 0\<close> have \<open>m div n mod q < q\<close> |
|
1287 |
by (rule mod_less_divisor) |
|
1288 |
then obtain s where \<open>q = Suc (m div n mod q + s)\<close> |
|
1289 |
by (blast dest: less_imp_Suc_add) |
|
1290 |
moreover have \<open>m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)\<close> |
|
1291 |
using \<open>m mod n < n\<close> by (simp add: add_mult_distrib2) |
|
1292 |
ultimately have \<open>m mod n + n * (m div n mod q) < n * q\<close> |
|
1293 |
by simp |
|
1294 |
then show ?case |
|
1295 |
by (simp add: algebra_simps flip: add_mult_distrib2) |
|
1296 |
qed |
|
1297 |
then show ?Q and ?R |
|
1298 |
by simp_all |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1299 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1300 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1301 |
lemma div_le_mono: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1302 |
"m div k \<le> n div k" if "m \<le> n" for m n k :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1303 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1304 |
from that obtain q where "n = m + q" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1305 |
by (auto simp add: le_iff_add) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1306 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1307 |
by (simp add: div_add1_eq [of m q k]) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1308 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1309 |
|
69593 | 1310 |
text \<open>Antimonotonicity of \<^const>\<open>divide\<close> in second argument\<close> |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1311 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1312 |
lemma div_le_mono2: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1313 |
"k div n \<le> k div m" if "0 < m" and "m \<le> n" for m n k :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1314 |
using that proof (induct k arbitrary: m rule: less_induct) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1315 |
case (less k) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1316 |
show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1317 |
proof (cases "n \<le> k") |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1318 |
case False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1319 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1320 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1321 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1322 |
case True |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1323 |
have "(k - n) div n \<le> (k - m) div n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1324 |
using less.prems |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1325 |
by (blast intro: div_le_mono diff_le_mono2) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1326 |
also have "\<dots> \<le> (k - m) div m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1327 |
using \<open>n \<le> k\<close> less.prems less.hyps [of "k - m" m] |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1328 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1329 |
finally show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1330 |
using \<open>n \<le> k\<close> less.prems |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1331 |
by (simp add: le_div_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1332 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1333 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1334 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1335 |
lemma div_le_dividend [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1336 |
"m div n \<le> m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1337 |
using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1338 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1339 |
lemma div_less_dividend [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1340 |
"m div n < m" if "1 < n" and "0 < m" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1341 |
using that proof (induct m rule: less_induct) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1342 |
case (less m) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1343 |
show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1344 |
proof (cases "n < m") |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1345 |
case False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1346 |
with less show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1347 |
by (cases "n = m") simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1348 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1349 |
case True |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1350 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1351 |
using less.hyps [of "m - n"] less.prems |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1352 |
by (simp add: le_div_geq) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1353 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1354 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1355 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1356 |
lemma div_eq_dividend_iff: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1357 |
"m div n = m \<longleftrightarrow> n = 1" if "m > 0" for m n :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1358 |
proof |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1359 |
assume "n = 1" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1360 |
then show "m div n = m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1361 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1362 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1363 |
assume P: "m div n = m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1364 |
show "n = 1" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1365 |
proof (rule ccontr) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1366 |
have "n \<noteq> 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1367 |
by (rule ccontr) (use that P in auto) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1368 |
moreover assume "n \<noteq> 1" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1369 |
ultimately have "n > 1" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1370 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1371 |
with that have "m div n < m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1372 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1373 |
with P show False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1374 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1375 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1376 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1377 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1378 |
lemma less_mult_imp_div_less: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1379 |
"m div n < i" if "m < i * n" for m n i :: nat |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1380 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1381 |
from that have "i * n > 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1382 |
by (cases "i * n = 0") simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1383 |
then have "i > 0" and "n > 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1384 |
by simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1385 |
have "m div n * n \<le> m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1386 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1387 |
then have "m div n * n < i * n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1388 |
using that by (rule le_less_trans) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1389 |
with \<open>n > 0\<close> show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1390 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1391 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1392 |
|
73853 | 1393 |
lemma div_less_iff_less_mult: |
1394 |
\<open>m div q < n \<longleftrightarrow> m < n * q\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
|
1395 |
if \<open>q > 0\<close> for m n q :: nat |
|
1396 |
proof |
|
1397 |
assume ?Q then show ?P |
|
1398 |
by (rule less_mult_imp_div_less) |
|
1399 |
next |
|
1400 |
assume ?P |
|
1401 |
then obtain h where \<open>n = Suc (m div q + h)\<close> |
|
1402 |
using less_natE by blast |
|
1403 |
moreover have \<open>m < m + (Suc h * q - m mod q)\<close> |
|
1404 |
using that by (simp add: trans_less_add1) |
|
1405 |
ultimately show ?Q |
|
1406 |
by (simp add: algebra_simps flip: minus_mod_eq_mult_div) |
|
1407 |
qed |
|
1408 |
||
1409 |
lemma less_eq_div_iff_mult_less_eq: |
|
1410 |
\<open>m \<le> n div q \<longleftrightarrow> m * q \<le> n\<close> if \<open>q > 0\<close> for m n q :: nat |
|
1411 |
using div_less_iff_less_mult [of q n m] that by auto |
|
1412 |
||
76231 | 1413 |
lemma div_Suc: |
76246 | 1414 |
\<open>Suc m div n = (if Suc m mod n = 0 then Suc (m div n) else m div n)\<close> |
76231 | 1415 |
proof (cases \<open>n = 0 \<or> n = 1\<close>) |
1416 |
case True |
|
1417 |
then show ?thesis by auto |
|
1418 |
next |
|
1419 |
case False |
|
1420 |
then have \<open>n > 1\<close> |
|
1421 |
by simp |
|
76246 | 1422 |
then have \<open>Suc m div n = m div n + Suc (m mod n) div n\<close> |
1423 |
using div_add1_eq [of m 1 n] by simp |
|
1424 |
also have \<open>Suc (m mod n) div n = of_bool (n dvd Suc m)\<close> |
|
76231 | 1425 |
proof (cases \<open>n dvd Suc m\<close>) |
1426 |
case False |
|
76246 | 1427 |
moreover have \<open>Suc (m mod n) \<noteq> n\<close> |
76231 | 1428 |
proof (rule ccontr) |
1429 |
assume \<open>\<not> Suc (m mod n) \<noteq> n\<close> |
|
76246 | 1430 |
then have \<open>m mod n = n - Suc 0\<close> |
76231 | 1431 |
by simp |
1432 |
with \<open>n > 1\<close> have \<open>(m + 1) mod n = 0\<close> |
|
1433 |
by (subst mod_add_left_eq [symmetric]) simp |
|
1434 |
then have \<open>n dvd Suc m\<close> |
|
1435 |
by auto |
|
1436 |
with False show False .. |
|
1437 |
qed |
|
1438 |
moreover have \<open>Suc (m mod n) \<le> n\<close> |
|
1439 |
using \<open>n > 1\<close> by (simp add: Suc_le_eq) |
|
76246 | 1440 |
ultimately show ?thesis |
1441 |
by (simp add: div_eq_0_iff) |
|
1442 |
next |
|
1443 |
case True |
|
1444 |
then obtain q where q: \<open>Suc m = n * q\<close> .. |
|
1445 |
moreover have \<open>q > 0\<close> by (rule ccontr) |
|
1446 |
(use q in simp) |
|
1447 |
ultimately have \<open>m mod n = n - Suc 0\<close> |
|
1448 |
using \<open>n > 1\<close> mult_le_cancel1 [of n \<open>Suc 0\<close> q] |
|
1449 |
by (auto intro: mod_nat_eqI) |
|
1450 |
with True \<open>n > 1\<close> show ?thesis |
|
76231 | 1451 |
by simp |
1452 |
qed |
|
76246 | 1453 |
finally show ?thesis |
1454 |
by (simp add: mod_greater_zero_iff_not_dvd) |
|
76231 | 1455 |
qed |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1456 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1457 |
lemma mod_Suc: |
76246 | 1458 |
\<open>Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))\<close> |
1459 |
proof (cases \<open>n = 0\<close>) |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1460 |
case True |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1461 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1462 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1463 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1464 |
case False |
76246 | 1465 |
moreover have \<open>Suc m mod n = Suc (m mod n) mod n\<close> |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1466 |
by (simp add: mod_simps) |
76246 | 1467 |
ultimately show ?thesis |
1468 |
by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq) |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1469 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1470 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1471 |
lemma Suc_times_mod_eq: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1472 |
"Suc (m * n) mod m = 1" if "Suc 0 < m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1473 |
using that by (simp add: mod_Suc) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1474 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1475 |
lemma Suc_times_numeral_mod_eq [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1476 |
"Suc (numeral k * n) mod numeral k = 1" if "numeral k \<noteq> (1::nat)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1477 |
by (rule Suc_times_mod_eq) (use that in simp) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1478 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1479 |
lemma Suc_div_le_mono [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1480 |
"m div n \<le> Suc m div n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1481 |
by (simp add: div_le_mono) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1482 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1483 |
text \<open>These lemmas collapse some needless occurrences of Suc: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1484 |
at least three Sucs, since two and fewer are rewritten back to Suc again! |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1485 |
We already have some rules to simplify operands smaller than 3.\<close> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1486 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1487 |
lemma div_Suc_eq_div_add3 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1488 |
"m div Suc (Suc (Suc n)) = m div (3 + n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1489 |
by (simp add: Suc3_eq_add_3) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1490 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1491 |
lemma mod_Suc_eq_mod_add3 [simp]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1492 |
"m mod Suc (Suc (Suc n)) = m mod (3 + n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1493 |
by (simp add: Suc3_eq_add_3) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1494 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1495 |
lemma Suc_div_eq_add3_div: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1496 |
"Suc (Suc (Suc m)) div n = (3 + m) div n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1497 |
by (simp add: Suc3_eq_add_3) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1498 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1499 |
lemma Suc_mod_eq_add3_mod: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1500 |
"Suc (Suc (Suc m)) mod n = (3 + m) mod n" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1501 |
by (simp add: Suc3_eq_add_3) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1502 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1503 |
lemmas Suc_div_eq_add3_div_numeral [simp] = |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1504 |
Suc_div_eq_add3_div [of _ "numeral v"] for v |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1505 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1506 |
lemmas Suc_mod_eq_add3_mod_numeral [simp] = |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1507 |
Suc_mod_eq_add3_mod [of _ "numeral v"] for v |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1508 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1509 |
lemma (in field_char_0) of_nat_div: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1510 |
"of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1511 |
proof - |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1512 |
have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1513 |
unfolding of_nat_add by (cases "n = 0") simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1514 |
then show ?thesis |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1515 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1516 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1517 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1518 |
text \<open>An ``induction'' law for modulus arithmetic.\<close> |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1519 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1520 |
lemma mod_induct [consumes 3, case_names step]: |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1521 |
"P m" if "P n" and "n < p" and "m < p" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1522 |
and step: "\<And>n. n < p \<Longrightarrow> P n \<Longrightarrow> P (Suc n mod p)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1523 |
using \<open>m < p\<close> proof (induct m) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1524 |
case 0 |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1525 |
show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1526 |
proof (rule ccontr) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1527 |
assume "\<not> P 0" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1528 |
from \<open>n < p\<close> have "0 < p" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1529 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1530 |
from \<open>n < p\<close> obtain m where "0 < m" and "p = n + m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1531 |
by (blast dest: less_imp_add_positive) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1532 |
with \<open>P n\<close> have "P (p - m)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1533 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1534 |
moreover have "\<not> P (p - m)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1535 |
using \<open>0 < m\<close> proof (induct m) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1536 |
case 0 |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1537 |
then show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1538 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1539 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1540 |
case (Suc m) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1541 |
show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1542 |
proof |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1543 |
assume P: "P (p - Suc m)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1544 |
with \<open>\<not> P 0\<close> have "Suc m < p" |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1545 |
by (auto intro: ccontr) |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1546 |
then have "Suc (p - Suc m) = p - m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1547 |
by arith |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1548 |
moreover from \<open>0 < p\<close> have "p - Suc m < p" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1549 |
by arith |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1550 |
with P step have "P ((Suc (p - Suc m)) mod p)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1551 |
by blast |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1552 |
ultimately show False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1553 |
using \<open>\<not> P 0\<close> Suc.hyps by (cases "m = 0") simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1554 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1555 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1556 |
ultimately show False |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1557 |
by blast |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1558 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1559 |
next |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1560 |
case (Suc m) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1561 |
then have "m < p" and mod: "Suc m mod p = Suc m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1562 |
by simp_all |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1563 |
from \<open>m < p\<close> have "P m" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1564 |
by (rule Suc.hyps) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1565 |
with \<open>m < p\<close> have "P (Suc m mod p)" |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1566 |
by (rule step) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1567 |
with mod show ?case |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1568 |
by simp |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1569 |
qed |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
1570 |
|
73555 | 1571 |
lemma funpow_mod_eq: \<^marker>\<open>contributor \<open>Lars Noschinski\<close>\<close> |
1572 |
\<open>(f ^^ (m mod n)) x = (f ^^ m) x\<close> if \<open>(f ^^ n) x = x\<close> |
|
1573 |
proof - |
|
1574 |
have \<open>(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\<close> |
|
1575 |
by simp |
|
1576 |
also have \<open>\<dots> = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\<close> |
|
1577 |
by (simp only: funpow_add funpow_mult ac_simps) simp |
|
1578 |
also have \<open>((f ^^ n) ^^ q) x = x\<close> for q |
|
1579 |
by (induction q) (use \<open>(f ^^ n) x = x\<close> in simp_all) |
|
1580 |
finally show ?thesis |
|
1581 |
by simp |
|
1582 |
qed |
|
1583 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1584 |
lemma mod_eq_dvd_iff_nat: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1585 |
\<open>m mod q = n mod q \<longleftrightarrow> q dvd m - n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1586 |
if \<open>m \<ge> n\<close> for m n q :: nat |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1587 |
proof |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1588 |
assume ?Q |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1589 |
then obtain s where \<open>m - n = q * s\<close> .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1590 |
with that have \<open>m = q * s + n\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1591 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1592 |
then show ?P |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1593 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1594 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1595 |
assume ?P |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1596 |
have \<open>m - n = m div q * q + m mod q - (n div q * q + n mod q)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1597 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1598 |
also have \<open>\<dots> = q * (m div q - n div q)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1599 |
by (simp only: algebra_simps \<open>?P\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1600 |
finally show ?Q .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1601 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1602 |
|
76231 | 1603 |
lemma mod_eq_iff_dvd_symdiff_nat: |
1604 |
\<open>m mod q = n mod q \<longleftrightarrow> q dvd nat \<bar>int m - int n\<bar>\<close> |
|
1605 |
by (auto simp add: abs_if mod_eq_dvd_iff_nat nat_diff_distrib dest: sym intro: sym) |
|
1606 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1607 |
lemma mod_eq_nat1E: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1608 |
fixes m n q :: nat |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1609 |
assumes "m mod q = n mod q" and "m \<ge> n" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1610 |
obtains s where "m = n + q * s" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1611 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1612 |
from assms have "q dvd m - n" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1613 |
by (simp add: mod_eq_dvd_iff_nat) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1614 |
then obtain s where "m - n = q * s" .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1615 |
with \<open>m \<ge> n\<close> have "m = n + q * s" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1616 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1617 |
with that show thesis . |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1618 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1619 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1620 |
lemma mod_eq_nat2E: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1621 |
fixes m n q :: nat |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1622 |
assumes "m mod q = n mod q" and "n \<ge> m" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1623 |
obtains s where "n = m + q * s" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1624 |
using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1625 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1626 |
lemma nat_mod_eq_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1627 |
"(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" (is "?lhs = ?rhs") |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1628 |
proof |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1629 |
assume H: "x mod n = y mod n" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1630 |
{ assume xy: "x \<le> y" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1631 |
from H have th: "y mod n = x mod n" by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1632 |
from mod_eq_nat1E [OF th xy] obtain q where "y = x + n * q" . |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1633 |
then have "x + n * q = y + n * 0" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1634 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1635 |
then have "\<exists>q1 q2. x + n * q1 = y + n * q2" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1636 |
by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1637 |
} |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1638 |
moreover |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1639 |
{ assume xy: "y \<le> x" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1640 |
from mod_eq_nat1E [OF H xy] obtain q where "x = y + n * q" . |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1641 |
then have "x + n * 0 = y + n * q" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1642 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1643 |
then have "\<exists>q1 q2. x + n * q1 = y + n * q2" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1644 |
by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1645 |
} |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1646 |
ultimately show ?rhs using linear[of x y] by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1647 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1648 |
assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1649 |
hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1650 |
thus ?lhs by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1651 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1652 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1653 |
|
76387 | 1654 |
subsection \<open>Division on \<^typ>\<open>int\<close>\<close> |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1655 |
|
77812 | 1656 |
text \<open> |
1657 |
The following specification of integer division rounds towards minus infinity |
|
1658 |
and is advocated by Donald Knuth. See \cite{leijen01} for an overview and |
|
1659 |
terminology of different possibilities to specify integer division; |
|
1660 |
there division rounding towards minus infinitiy is named ``F-division''. |
|
1661 |
\<close> |
|
1662 |
||
75876 | 1663 |
subsubsection \<open>Basic instantiation\<close> |
1664 |
||
1665 |
instantiation int :: "{normalization_semidom, idom_modulo}" |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1666 |
begin |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1667 |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1668 |
definition normalize_int :: \<open>int \<Rightarrow> int\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1669 |
where [simp]: \<open>normalize = (abs :: int \<Rightarrow> int)\<close> |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1670 |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1671 |
definition unit_factor_int :: \<open>int \<Rightarrow> int\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1672 |
where [simp]: \<open>unit_factor = (sgn :: int \<Rightarrow> int)\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1673 |
|
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1674 |
definition divide_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1675 |
where \<open>k div l = (sgn k * sgn l * int (nat \<bar>k\<bar> div nat \<bar>l\<bar>) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1676 |
- of_bool (l \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close> |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1677 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1678 |
lemma divide_int_unfold: |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1679 |
\<open>(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1680 |
- of_bool ((k = 0 \<longleftrightarrow> m = 0) \<and> l \<noteq> 0 \<and> n \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> n dvd m))\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1681 |
by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1682 |
|
75876 | 1683 |
definition modulo_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
1684 |
where \<open>k mod l = sgn k * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>) + l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close> |
|
1685 |
||
1686 |
lemma modulo_int_unfold: |
|
1687 |
\<open>(sgn k * int m) mod (sgn l * int n) = |
|
1688 |
sgn k * int (m mod (of_bool (l \<noteq> 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \<longleftrightarrow> m = 0) \<and> sgn k \<noteq> sgn l \<and> \<not> n dvd m)\<close> |
|
1689 |
by (auto simp add: modulo_int_def sgn_mult abs_mult) |
|
1690 |
||
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1691 |
instance proof |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1692 |
fix k :: int show "k div 0 = 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1693 |
by (simp add: divide_int_def) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1694 |
next |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1695 |
fix k l :: int |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1696 |
assume "l \<noteq> 0" |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1697 |
obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1698 |
by (blast intro: int_sgnE elim: that) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1699 |
then have "k * l = sgn (s * t) * int (n * m)" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1700 |
by (simp add: ac_simps sgn_mult) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1701 |
with k l \<open>l \<noteq> 0\<close> show "k * l div l = k" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1702 |
by (simp only: divide_int_unfold) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1703 |
(auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) |
75876 | 1704 |
next |
1705 |
fix k l :: int |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1706 |
obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" |
75876 | 1707 |
by (blast intro: int_sgnE elim: that) |
1708 |
then show "k div l * l + k mod l = k" |
|
1709 |
by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff) |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1710 |
qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1711 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1712 |
end |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1713 |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1714 |
|
75876 | 1715 |
subsubsection \<open>Algebraic foundations\<close> |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1716 |
|
67051 | 1717 |
lemma coprime_int_iff [simp]: |
1718 |
"coprime (int m) (int n) \<longleftrightarrow> coprime m n" (is "?P \<longleftrightarrow> ?Q") |
|
1719 |
proof |
|
1720 |
assume ?P |
|
1721 |
show ?Q |
|
1722 |
proof (rule coprimeI) |
|
1723 |
fix q |
|
1724 |
assume "q dvd m" "q dvd n" |
|
1725 |
then have "int q dvd int m" "int q dvd int n" |
|
67118 | 1726 |
by simp_all |
67051 | 1727 |
with \<open>?P\<close> have "is_unit (int q)" |
1728 |
by (rule coprime_common_divisor) |
|
1729 |
then show "is_unit q" |
|
1730 |
by simp |
|
1731 |
qed |
|
1732 |
next |
|
1733 |
assume ?Q |
|
1734 |
show ?P |
|
1735 |
proof (rule coprimeI) |
|
1736 |
fix k |
|
1737 |
assume "k dvd int m" "k dvd int n" |
|
1738 |
then have "nat \<bar>k\<bar> dvd m" "nat \<bar>k\<bar> dvd n" |
|
67118 | 1739 |
by simp_all |
67051 | 1740 |
with \<open>?Q\<close> have "is_unit (nat \<bar>k\<bar>)" |
1741 |
by (rule coprime_common_divisor) |
|
1742 |
then show "is_unit k" |
|
1743 |
by simp |
|
1744 |
qed |
|
1745 |
qed |
|
1746 |
||
1747 |
lemma coprime_abs_left_iff [simp]: |
|
1748 |
"coprime \<bar>k\<bar> l \<longleftrightarrow> coprime k l" for k l :: int |
|
1749 |
using coprime_normalize_left_iff [of k l] by simp |
|
1750 |
||
1751 |
lemma coprime_abs_right_iff [simp]: |
|
1752 |
"coprime k \<bar>l\<bar> \<longleftrightarrow> coprime k l" for k l :: int |
|
1753 |
using coprime_abs_left_iff [of l k] by (simp add: ac_simps) |
|
1754 |
||
1755 |
lemma coprime_nat_abs_left_iff [simp]: |
|
1756 |
"coprime (nat \<bar>k\<bar>) n \<longleftrightarrow> coprime k (int n)" |
|
1757 |
proof - |
|
1758 |
define m where "m = nat \<bar>k\<bar>" |
|
1759 |
then have "\<bar>k\<bar> = int m" |
|
1760 |
by simp |
|
1761 |
moreover have "coprime k (int n) \<longleftrightarrow> coprime \<bar>k\<bar> (int n)" |
|
1762 |
by simp |
|
1763 |
ultimately show ?thesis |
|
1764 |
by simp |
|
1765 |
qed |
|
1766 |
||
1767 |
lemma coprime_nat_abs_right_iff [simp]: |
|
1768 |
"coprime n (nat \<bar>k\<bar>) \<longleftrightarrow> coprime (int n) k" |
|
1769 |
using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) |
|
1770 |
||
1771 |
lemma coprime_common_divisor_int: "coprime a b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> \<bar>x\<bar> = 1" |
|
1772 |
for a b :: int |
|
1773 |
by (drule coprime_common_divisor [of _ _ x]) simp_all |
|
1774 |
||
75876 | 1775 |
|
1776 |
subsubsection \<open>Basic conversions\<close> |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1777 |
|
75876 | 1778 |
lemma div_abs_eq_div_nat: |
1779 |
"\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)" |
|
1780 |
by (auto simp add: divide_int_def) |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1781 |
|
75876 | 1782 |
lemma div_eq_div_abs: |
1783 |
\<open>k div l = sgn k * sgn l * (\<bar>k\<bar> div \<bar>l\<bar>) |
|
1784 |
- of_bool (l \<noteq> 0 \<and> sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close> |
|
1785 |
for k l :: int |
|
1786 |
by (simp add: divide_int_def [of k l] div_abs_eq_div_nat) |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1787 |
|
75876 | 1788 |
lemma div_abs_eq: |
1789 |
\<open>\<bar>k\<bar> div \<bar>l\<bar> = sgn k * sgn l * (k div l + of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close> |
|
1790 |
for k l :: int |
|
1791 |
by (simp add: div_eq_div_abs [of k l] ac_simps) |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1792 |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1793 |
lemma mod_abs_eq_div_nat: |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1794 |
"\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)" |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1795 |
by (simp add: modulo_int_def) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1796 |
|
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1797 |
lemma mod_eq_mod_abs: |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1798 |
\<open>k mod l = sgn k * (\<bar>k\<bar> mod \<bar>l\<bar>) + l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k)\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1799 |
for k l :: int |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1800 |
by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1801 |
|
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1802 |
lemma mod_abs_eq: |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1803 |
\<open>\<bar>k\<bar> mod \<bar>l\<bar> = sgn k * (k mod l - l * of_bool (sgn k \<noteq> sgn l \<and> \<not> l dvd k))\<close> |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1804 |
for k l :: int |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1805 |
by (auto simp: mod_eq_mod_abs [of k l]) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1806 |
|
75876 | 1807 |
lemma div_sgn_abs_cancel: |
1808 |
fixes k l v :: int |
|
1809 |
assumes "v \<noteq> 0" |
|
1810 |
shows "(sgn v * \<bar>k\<bar>) div (sgn v * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>" |
|
1811 |
using assms by (simp add: sgn_mult abs_mult sgn_0_0 |
|
1812 |
divide_int_def [of "sgn v * \<bar>k\<bar>" "sgn v * \<bar>l\<bar>"] flip: div_abs_eq_div_nat) |
|
1813 |
||
1814 |
lemma div_eq_sgn_abs: |
|
1815 |
fixes k l v :: int |
|
1816 |
assumes "sgn k = sgn l" |
|
1817 |
shows "k div l = \<bar>k\<bar> div \<bar>l\<bar>" |
|
1818 |
using assms by (auto simp add: div_abs_eq) |
|
1819 |
||
1820 |
lemma div_dvd_sgn_abs: |
|
1821 |
fixes k l :: int |
|
1822 |
assumes "l dvd k" |
|
1823 |
shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)" |
|
1824 |
using assms by (auto simp add: div_abs_eq ac_simps) |
|
1825 |
||
1826 |
lemma div_noneq_sgn_abs: |
|
1827 |
fixes k l :: int |
|
1828 |
assumes "l \<noteq> 0" |
|
1829 |
assumes "sgn k \<noteq> sgn l" |
|
1830 |
shows "k div l = - (\<bar>k\<bar> div \<bar>l\<bar>) - of_bool (\<not> l dvd k)" |
|
1831 |
using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp) |
|
1832 |
||
1833 |
||
1834 |
subsubsection \<open>Euclidean division\<close> |
|
1835 |
||
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1836 |
instantiation int :: unique_euclidean_ring |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1837 |
begin |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1838 |
|
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1839 |
definition euclidean_size_int :: "int \<Rightarrow> nat" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1840 |
where [simp]: "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1841 |
|
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1842 |
definition division_segment_int :: "int \<Rightarrow> int" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1843 |
where "division_segment_int k = (if k \<ge> 0 then 1 else - 1)" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1844 |
|
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1845 |
lemma division_segment_eq_sgn: |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1846 |
"division_segment k = sgn k" if "k \<noteq> 0" for k :: int |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1847 |
using that by (simp add: division_segment_int_def) |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1848 |
|
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1849 |
lemma abs_division_segment [simp]: |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1850 |
"\<bar>division_segment k\<bar> = 1" for k :: int |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1851 |
by (simp add: division_segment_int_def) |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1852 |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1853 |
lemma abs_mod_less: |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1854 |
"\<bar>k mod l\<bar> < \<bar>l\<bar>" if "l \<noteq> 0" for k l :: int |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1855 |
proof - |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1856 |
obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1857 |
by (blast intro: int_sgnE elim: that) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1858 |
with that show ?thesis |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1859 |
by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1860 |
simp flip: right_diff_distrib dest!: sgn_not_eq_imp) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1861 |
(simp add: sgn_0_0) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1862 |
qed |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1863 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1864 |
lemma sgn_mod: |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1865 |
"sgn (k mod l) = sgn l" if "l \<noteq> 0" "\<not> l dvd k" for k l :: int |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1866 |
proof - |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1867 |
obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1868 |
by (blast intro: int_sgnE elim: that) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1869 |
with that show ?thesis |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1870 |
by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1871 |
simp flip: right_diff_distrib dest!: sgn_not_eq_imp) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1872 |
qed |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1873 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1874 |
instance proof |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1875 |
fix k l :: int |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1876 |
show "division_segment (k mod l) = division_segment l" if |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1877 |
"l \<noteq> 0" and "\<not> l dvd k" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1878 |
using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1879 |
next |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1880 |
fix l q r :: int |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1881 |
obtain n m and s t |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1882 |
where l: "l = sgn s * int n" and q: "q = sgn t * int m" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1883 |
by (blast intro: int_sgnE elim: that) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1884 |
assume \<open>l \<noteq> 0\<close> |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1885 |
with l have "s \<noteq> 0" and "n > 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1886 |
by (simp_all add: sgn_0_0) |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1887 |
assume "division_segment r = division_segment l" |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1888 |
moreover have "r = sgn r * \<bar>r\<bar>" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1889 |
by (simp add: sgn_mult_abs) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1890 |
moreover define u where "u = nat \<bar>r\<bar>" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1891 |
ultimately have "r = sgn l * int u" |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66837
diff
changeset
|
1892 |
using division_segment_eq_sgn \<open>l \<noteq> 0\<close> by (cases "r = 0") simp_all |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1893 |
with l \<open>n > 0\<close> have r: "r = sgn s * int u" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1894 |
by (simp add: sgn_mult) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1895 |
assume "euclidean_size r < euclidean_size l" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1896 |
with l r \<open>s \<noteq> 0\<close> have "u < n" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1897 |
by (simp add: abs_mult) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1898 |
show "(q * l + r) div l = q" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1899 |
proof (cases "q = 0 \<or> r = 0") |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1900 |
case True |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1901 |
then show ?thesis |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1902 |
proof |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1903 |
assume "q = 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1904 |
then show ?thesis |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1905 |
using l r \<open>u < n\<close> by (simp add: divide_int_unfold) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1906 |
next |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1907 |
assume "r = 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1908 |
from \<open>r = 0\<close> have *: "q * l + r = sgn (t * s) * int (n * m)" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1909 |
using q l by (simp add: ac_simps sgn_mult) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1910 |
from \<open>s \<noteq> 0\<close> \<open>n > 0\<close> show ?thesis |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1911 |
by (simp only: *, simp only: * q l divide_int_unfold) |
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
1912 |
(auto simp add: sgn_mult ac_simps) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1913 |
qed |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1914 |
next |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1915 |
case False |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1916 |
with q r have "t \<noteq> 0" and "m > 0" and "s \<noteq> 0" and "u > 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1917 |
by (simp_all add: sgn_0_0) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1918 |
moreover from \<open>0 < m\<close> \<open>u < n\<close> have "u \<le> m * n" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1919 |
using mult_le_less_imp_less [of 1 m u n] by simp |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1920 |
ultimately have *: "q * l + r = sgn (s * t) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1921 |
* int (if t < 0 then m * n - u else m * n + u)" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1922 |
using l q r |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1923 |
by (simp add: sgn_mult algebra_simps of_nat_diff) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1924 |
have "(m * n - u) div n = m - 1" if "u > 0" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1925 |
using \<open>0 < m\<close> \<open>u < n\<close> that |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1926 |
by (auto intro: div_nat_eqI simp add: algebra_simps) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1927 |
moreover have "n dvd m * n - u \<longleftrightarrow> n dvd u" |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1928 |
using \<open>u \<le> m * n\<close> dvd_diffD1 [of n "m * n" u] |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1929 |
by auto |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1930 |
ultimately show ?thesis |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1931 |
using \<open>s \<noteq> 0\<close> \<open>m > 0\<close> \<open>u > 0\<close> \<open>u < n\<close> \<open>u \<le> m * n\<close> |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1932 |
by (simp only: *, simp only: l q divide_int_unfold) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1933 |
(auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1934 |
qed |
68536 | 1935 |
qed (use mult_le_mono2 [of 1] in \<open>auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\<close>) |
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1936 |
|
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1937 |
end |
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1938 |
|
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1939 |
lemma euclidean_relation_intI [case_names by0 divides euclidean_relation]: |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1940 |
\<open>(k div l, k mod l) = (q, r)\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1941 |
if by0': \<open>l = 0 \<Longrightarrow> q = 0 \<and> r = k\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1942 |
and divides': \<open>l \<noteq> 0 \<Longrightarrow> l dvd k \<Longrightarrow> r = 0 \<and> k = q * l\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1943 |
and euclidean_relation': \<open>l \<noteq> 0 \<Longrightarrow> \<not> l dvd k \<Longrightarrow> sgn r = sgn l |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1944 |
\<and> \<bar>r\<bar> < \<bar>l\<bar> \<and> k = q * l + r\<close> for k l :: int |
76245 | 1945 |
proof (induction rule: euclidean_relationI) |
76141
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1946 |
case by0 |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1947 |
then show ?case |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1948 |
by (rule by0') |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1949 |
next |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1950 |
case divides |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1951 |
then show ?case |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1952 |
by (rule divides') |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1953 |
next |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1954 |
case euclidean_relation |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1955 |
with euclidean_relation' have \<open>sgn r = sgn l\<close> \<open>\<bar>r\<bar> < \<bar>l\<bar>\<close> \<open>k = q * l + r\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1956 |
by simp_all |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1957 |
from \<open>sgn r = sgn l\<close> \<open>l \<noteq> 0\<close> have \<open>division_segment r = division_segment l\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1958 |
by (simp add: division_segment_int_def sgn_if split: if_splits) |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1959 |
with \<open>\<bar>r\<bar> < \<bar>l\<bar>\<close> \<open>k = q * l + r\<close> |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1960 |
show ?case |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1961 |
by simp |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1962 |
qed |
e7497a1de8b9
more concise instance-specific rules on euclidean relation
haftmann
parents:
76053
diff
changeset
|
1963 |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66814
diff
changeset
|
1964 |
|
75876 | 1965 |
subsubsection \<open>Trivial reduction steps\<close> |
1966 |
||
1967 |
lemma div_pos_pos_trivial [simp]: |
|
1968 |
"k div l = 0" if "k \<ge> 0" and "k < l" for k l :: int |
|
1969 |
using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) |
|
1970 |
||
1971 |
lemma mod_pos_pos_trivial [simp]: |
|
1972 |
"k mod l = k" if "k \<ge> 0" and "k < l" for k l :: int |
|
1973 |
using that by (simp add: mod_eq_self_iff_div_eq_0) |
|
1974 |
||
1975 |
lemma div_neg_neg_trivial [simp]: |
|
1976 |
"k div l = 0" if "k \<le> 0" and "l < k" for k l :: int |
|
1977 |
using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) |
|
1978 |
||
1979 |
lemma mod_neg_neg_trivial [simp]: |
|
1980 |
"k mod l = k" if "k \<le> 0" and "l < k" for k l :: int |
|
1981 |
using that by (simp add: mod_eq_self_iff_div_eq_0) |
|
1982 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1983 |
lemma |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1984 |
div_pos_neg_trivial: \<open>k div l = - 1\<close> (is ?Q) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1985 |
and mod_pos_neg_trivial: \<open>k mod l = k + l\<close> (is ?R) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1986 |
if \<open>0 < k\<close> and \<open>k + l \<le> 0\<close> for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1987 |
proof - |
76053 | 1988 |
from that have \<open>l < 0\<close> |
1989 |
by simp |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1990 |
have \<open>(k div l, k mod l) = (- 1, k + l)\<close> |
76245 | 1991 |
proof (induction rule: euclidean_relation_intI) |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1992 |
case by0 |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1993 |
with \<open>l < 0\<close> show ?case |
76053 | 1994 |
by simp |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1995 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1996 |
case divides |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1997 |
from \<open>l dvd k\<close> obtain j where \<open>k = l * j\<close> .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1998 |
with \<open>l < 0\<close> \<open>0 < k\<close> have \<open>j < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
1999 |
by (simp add: zero_less_mult_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2000 |
moreover from \<open>k + l \<le> 0\<close> \<open>k = l * j\<close> have \<open>l * (j + 1) \<le> 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2001 |
by (simp add: algebra_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2002 |
with \<open>l < 0\<close> have \<open>j + 1 \<ge> 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2003 |
by (simp add: mult_le_0_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2004 |
with \<open>j < 0\<close> have \<open>j = - 1\<close> |
76053 | 2005 |
by simp |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2006 |
with \<open>k = l * j\<close> show ?case |
76053 | 2007 |
by simp |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2008 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2009 |
case euclidean_relation |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2010 |
with \<open>k + l \<le> 0\<close> have \<open>k + l < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2011 |
by (auto simp add: less_le add_eq_0_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2012 |
with \<open>0 < k\<close> show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2013 |
by simp |
76053 | 2014 |
qed |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2015 |
then show ?Q and ?R |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2016 |
by simp_all |
75876 | 2017 |
qed |
2018 |
||
2019 |
text \<open>There is neither \<open>div_neg_pos_trivial\<close> nor \<open>mod_neg_pos_trivial\<close> |
|
2020 |
because \<^term>\<open>0 div l = 0\<close> would supersede it.\<close> |
|
2021 |
||
2022 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2023 |
subsubsection \<open>More uniqueness rules\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2024 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2025 |
lemma |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2026 |
fixes a b q r :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2027 |
assumes \<open>a = b * q + r\<close> \<open>0 \<le> r\<close> \<open>r < b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2028 |
shows int_div_pos_eq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2029 |
\<open>a div b = q\<close> (is ?Q) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2030 |
and int_mod_pos_eq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2031 |
\<open>a mod b = r\<close> (is ?R) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2032 |
proof - |
76245 | 2033 |
have \<open>(a div b, a mod b) = (q, r)\<close> |
2034 |
by (induction rule: euclidean_relation_intI) |
|
2035 |
(use assms in \<open>auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le\<close>) |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2036 |
then show ?Q and ?R |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2037 |
by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2038 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2039 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2040 |
lemma int_div_neg_eq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2041 |
\<open>a div b = q\<close> if \<open>a = b * q + r\<close> \<open>r \<le> 0\<close> \<open>b < r\<close> for a b q r :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2042 |
using that int_div_pos_eq [of a \<open>- b\<close> \<open>- q\<close> \<open>- r\<close>] by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2043 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2044 |
lemma int_mod_neg_eq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2045 |
\<open>a mod b = r\<close> if \<open>a = b * q + r\<close> \<open>r \<le> 0\<close> \<open>b < r\<close> for a b q r :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2046 |
using that int_div_neg_eq [of a b q r] by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2047 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2048 |
|
75877 | 2049 |
subsubsection \<open>Laws for unary minus\<close> |
2050 |
||
2051 |
lemma zmod_zminus1_not_zero: |
|
2052 |
fixes k l :: int |
|
2053 |
shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" |
|
2054 |
by (simp add: mod_eq_0_iff_dvd) |
|
2055 |
||
2056 |
lemma zmod_zminus2_not_zero: |
|
2057 |
fixes k l :: int |
|
2058 |
shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" |
|
2059 |
by (simp add: mod_eq_0_iff_dvd) |
|
2060 |
||
2061 |
lemma zdiv_zminus1_eq_if: |
|
2062 |
\<open>(- a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\<close> |
|
2063 |
if \<open>b \<noteq> 0\<close> for a b :: int |
|
2064 |
using that sgn_not_eq_imp [of b \<open>- a\<close>] |
|
2065 |
by (cases \<open>a = 0\<close>) (auto simp add: div_eq_div_abs [of \<open>- a\<close> b] div_eq_div_abs [of a b] sgn_eq_0_iff) |
|
2066 |
||
2067 |
lemma zdiv_zminus2_eq_if: |
|
2068 |
\<open>a div (- b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\<close> |
|
2069 |
if \<open>b \<noteq> 0\<close> for a b :: int |
|
2070 |
using that by (auto simp add: zdiv_zminus1_eq_if div_minus_right) |
|
2071 |
||
2072 |
lemma zmod_zminus1_eq_if: |
|
2073 |
\<open>(- a) mod b = (if a mod b = 0 then 0 else b - (a mod b))\<close> |
|
2074 |
for a b :: int |
|
2075 |
by (cases \<open>b = 0\<close>) |
|
2076 |
(auto simp flip: minus_div_mult_eq_mod simp add: zdiv_zminus1_eq_if algebra_simps) |
|
2077 |
||
2078 |
lemma zmod_zminus2_eq_if: |
|
2079 |
\<open>a mod (- b) = (if a mod b = 0 then 0 else (a mod b) - b)\<close> |
|
2080 |
for a b :: int |
|
2081 |
by (auto simp add: zmod_zminus1_eq_if mod_minus_right) |
|
2082 |
||
2083 |
||
75876 | 2084 |
subsubsection \<open>Borders\<close> |
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2085 |
|
75876 | 2086 |
lemma pos_mod_bound [simp]: |
2087 |
"k mod l < l" if "l > 0" for k l :: int |
|
2088 |
proof - |
|
2089 |
obtain m and s where "k = sgn s * int m" |
|
2090 |
by (rule int_sgnE) |
|
2091 |
moreover from that obtain n where "l = sgn 1 * int n" |
|
2092 |
by (cases l) simp_all |
|
2093 |
moreover from this that have "n > 0" |
|
2094 |
by simp |
|
2095 |
ultimately show ?thesis |
|
2096 |
by (simp only: modulo_int_unfold) |
|
2097 |
(auto simp add: mod_greater_zero_iff_not_dvd sgn_1_pos) |
|
2098 |
qed |
|
2099 |
||
2100 |
lemma neg_mod_bound [simp]: |
|
2101 |
"l < k mod l" if "l < 0" for k l :: int |
|
2102 |
proof - |
|
2103 |
obtain m and s where "k = sgn s * int m" |
|
2104 |
by (rule int_sgnE) |
|
2105 |
moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" |
|
2106 |
by (cases l) simp_all |
|
2107 |
moreover define n where "n = Suc q" |
|
2108 |
then have "Suc q = n" |
|
2109 |
by simp |
|
2110 |
ultimately show ?thesis |
|
2111 |
by (simp only: modulo_int_unfold) |
|
2112 |
(auto simp add: mod_greater_zero_iff_not_dvd sgn_1_neg) |
|
2113 |
qed |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2114 |
|
75876 | 2115 |
lemma pos_mod_sign [simp]: |
2116 |
"0 \<le> k mod l" if "l > 0" for k l :: int |
|
2117 |
proof - |
|
2118 |
obtain m and s where "k = sgn s * int m" |
|
2119 |
by (rule int_sgnE) |
|
2120 |
moreover from that obtain n where "l = sgn 1 * int n" |
|
2121 |
by (cases l) auto |
|
2122 |
moreover from this that have "n > 0" |
|
2123 |
by simp |
|
2124 |
ultimately show ?thesis |
|
2125 |
by (simp only: modulo_int_unfold) (auto simp add: sgn_1_pos) |
|
2126 |
qed |
|
2127 |
||
2128 |
lemma neg_mod_sign [simp]: |
|
2129 |
"k mod l \<le> 0" if "l < 0" for k l :: int |
|
2130 |
proof - |
|
2131 |
obtain m and s where "k = sgn s * int m" |
|
2132 |
by (rule int_sgnE) |
|
2133 |
moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" |
|
2134 |
by (cases l) simp_all |
|
2135 |
moreover define n where "n = Suc q" |
|
2136 |
then have "Suc q = n" |
|
2137 |
by simp |
|
2138 |
moreover have \<open>int (m mod n) \<le> int n\<close> |
|
2139 |
using \<open>Suc q = n\<close> by simp |
|
2140 |
then have \<open>sgn s * int (m mod n) \<le> int n\<close> |
|
2141 |
by (cases s \<open>0::int\<close> rule: linorder_cases) simp_all |
|
2142 |
ultimately show ?thesis |
|
2143 |
by (simp only: modulo_int_unfold) auto |
|
2144 |
qed |
|
2145 |
||
2146 |
||
75881 | 2147 |
subsubsection \<open>Splitting Rules for div and mod\<close> |
2148 |
||
2149 |
lemma split_zdiv: |
|
2150 |
\<open>P (n div k) \<longleftrightarrow> |
|
2151 |
(k = 0 \<longrightarrow> P 0) \<and> |
|
2152 |
(0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> P i)) \<and> |
|
2153 |
(k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> P i))\<close> (is ?div) |
|
2154 |
and split_zmod: |
|
2155 |
\<open>Q (n mod k) \<longleftrightarrow> |
|
2156 |
(k = 0 \<longrightarrow> Q n) \<and> |
|
2157 |
(0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> Q j)) \<and> |
|
2158 |
(k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> Q j))\<close> (is ?mod) |
|
2159 |
for n k :: int |
|
2160 |
proof - |
|
2161 |
have *: \<open>R (n div k) (n mod k) \<longleftrightarrow> |
|
2162 |
(k = 0 \<longrightarrow> R 0 n) \<and> |
|
2163 |
(0 < k \<longrightarrow> (\<forall>i j. 0 \<le> j \<and> j < k \<and> n = k * i + j \<longrightarrow> R i j)) \<and> |
|
2164 |
(k < 0 \<longrightarrow> (\<forall>i j. k < j \<and> j \<le> 0 \<and> n = k * i + j \<longrightarrow> R i j))\<close> for R |
|
2165 |
by (cases \<open>k = 0\<close>) |
|
2166 |
(auto simp add: linorder_class.neq_iff) |
|
2167 |
from * [of \<open>\<lambda>q _. P q\<close>] show ?div . |
|
2168 |
from * [of \<open>\<lambda>_ r. Q r\<close>] show ?mod . |
|
2169 |
qed |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2170 |
|
75881 | 2171 |
text \<open>Enable (lin)arith to deal with \<^const>\<open>divide\<close> and \<^const>\<open>modulo\<close> |
2172 |
when these are applied to some constant that is of the form |
|
2173 |
\<^term>\<open>numeral k\<close>:\<close> |
|
2174 |
declare split_zdiv [of _ _ \<open>numeral n\<close>, linarith_split] for n |
|
2175 |
declare split_zdiv [of _ _ \<open>- numeral n\<close>, linarith_split] for n |
|
2176 |
declare split_zmod [of _ _ \<open>numeral n\<close>, linarith_split] for n |
|
2177 |
declare split_zmod [of _ _ \<open>- numeral n\<close>, linarith_split] for n |
|
2178 |
||
2179 |
lemma zdiv_eq_0_iff: |
|
2180 |
"i div k = 0 \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" (is "?L = ?R") |
|
2181 |
for i k :: int |
|
2182 |
proof |
|
2183 |
assume ?L |
|
2184 |
moreover have "?L \<longrightarrow> ?R" |
|
2185 |
by (rule split_zdiv [THEN iffD2]) simp |
|
2186 |
ultimately show ?R |
|
2187 |
by blast |
|
2188 |
next |
|
2189 |
assume ?R then show ?L |
|
2190 |
by auto |
|
2191 |
qed |
|
2192 |
||
2193 |
lemma zmod_trivial_iff: |
|
2194 |
fixes i k :: int |
|
2195 |
shows "i mod k = i \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" |
|
2196 |
proof - |
|
2197 |
have "i mod k = i \<longleftrightarrow> i div k = 0" |
|
2198 |
using div_mult_mod_eq [of i k] by safe auto |
|
2199 |
with zdiv_eq_0_iff |
|
2200 |
show ?thesis |
|
2201 |
by simp |
|
2202 |
qed |
|
2203 |
||
2204 |
||
75876 | 2205 |
subsubsection \<open>Algebraic rewrites\<close> |
71157 | 2206 |
|
76387 | 2207 |
lemma zdiv_zmult2_eq: \<open>a div (b * c) = (a div b) div c\<close> (is ?Q) |
2208 |
and zmod_zmult2_eq: \<open>a mod (b * c) = b * (a div b mod c) + a mod b\<close> (is ?P) |
|
2209 |
if \<open>c \<ge> 0\<close> for a b c :: int |
|
2210 |
proof - |
|
2211 |
have *: \<open>(a div (b * c), a mod (b * c)) = ((a div b) div c, b * (a div b mod c) + a mod b)\<close> |
|
2212 |
if \<open>b > 0\<close> for a b |
|
2213 |
proof (induction rule: euclidean_relationI) |
|
2214 |
case by0 |
|
2215 |
then show ?case by auto |
|
2216 |
next |
|
2217 |
case divides |
|
2218 |
then obtain d where \<open>a = b * c * d\<close> |
|
2219 |
by blast |
|
2220 |
with divides that show ?case |
|
2221 |
by (simp add: ac_simps) |
|
2222 |
next |
|
2223 |
case euclidean_relation |
|
2224 |
with \<open>b > 0\<close> \<open>c \<ge> 0\<close> have \<open>0 < c\<close> \<open>b > 0\<close> |
|
2225 |
by simp_all |
|
2226 |
then have \<open>a mod b < b\<close> |
|
2227 |
by simp |
|
2228 |
moreover have \<open>1 \<le> c - a div b mod c\<close> |
|
2229 |
using \<open>c > 0\<close> by (simp add: int_one_le_iff_zero_less) |
|
2230 |
ultimately have \<open>a mod b * 1 < b * (c - a div b mod c)\<close> |
|
2231 |
by (rule mult_less_le_imp_less) (use \<open>b > 0\<close> in simp_all) |
|
2232 |
with \<open>0 < b\<close> \<open>0 < c\<close> show ?case |
|
2233 |
by (simp add: division_segment_int_def algebra_simps flip: minus_mod_eq_mult_div) |
|
2234 |
qed |
|
2235 |
show ?Q |
|
2236 |
proof (cases \<open>b \<ge> 0\<close>) |
|
2237 |
case True |
|
2238 |
with * [of b a] show ?thesis |
|
2239 |
by (cases \<open>b = 0\<close>) simp_all |
|
2240 |
next |
|
2241 |
case False |
|
2242 |
with * [of \<open>- b\<close> \<open>- a\<close>] show ?thesis |
|
2243 |
by simp |
|
2244 |
qed |
|
2245 |
show ?P |
|
2246 |
proof (cases \<open>b \<ge> 0\<close>) |
|
2247 |
case True |
|
2248 |
with * [of b a] show ?thesis |
|
2249 |
by (cases \<open>b = 0\<close>) simp_all |
|
2250 |
next |
|
2251 |
case False |
|
2252 |
with * [of \<open>- b\<close> \<open>- a\<close>] show ?thesis |
|
2253 |
by simp |
|
2254 |
qed |
|
74592 | 2255 |
qed |
2256 |
||
75876 | 2257 |
lemma zdiv_zmult2_eq': |
2258 |
\<open>k div (l * j) = ((sgn j * k) div l) div \<bar>j\<bar>\<close> for k l j :: int |
|
2259 |
proof - |
|
2260 |
have \<open>k div (l * j) = (sgn j * k) div (sgn j * (l * j))\<close> |
|
2261 |
by (simp add: sgn_0_0) |
|
2262 |
also have \<open>sgn j * (l * j) = l * \<bar>j\<bar>\<close> |
|
2263 |
by (simp add: mult.left_commute [of _ l] abs_sgn) (simp add: ac_simps) |
|
2264 |
also have \<open>(sgn j * k) div (l * \<bar>j\<bar>) = ((sgn j * k) div l) div \<bar>j\<bar>\<close> |
|
2265 |
by (simp add: zdiv_zmult2_eq) |
|
2266 |
finally show ?thesis . |
|
2267 |
qed |
|
2268 |
||
75881 | 2269 |
lemma half_nonnegative_int_iff [simp]: |
2270 |
\<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int |
|
2271 |
by auto |
|
2272 |
||
2273 |
lemma half_negative_int_iff [simp]: |
|
2274 |
\<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> for k :: int |
|
2275 |
by auto |
|
2276 |
||
75876 | 2277 |
|
2278 |
subsubsection \<open>Distributive laws for conversions.\<close> |
|
2279 |
||
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2280 |
lemma zdiv_int: |
76387 | 2281 |
\<open>int (m div n) = int m div int n\<close> |
2282 |
by (cases \<open>m = 0\<close>) (auto simp add: divide_int_def) |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2283 |
|
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2284 |
lemma zmod_int: |
76387 | 2285 |
\<open>int (m mod n) = int m mod int n\<close> |
2286 |
by (cases \<open>m = 0\<close>) (auto simp add: modulo_int_def) |
|
75875
48d032035744
streamlined primitive definitions for integer division
haftmann
parents:
75669
diff
changeset
|
2287 |
|
75876 | 2288 |
lemma nat_div_distrib: |
2289 |
\<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> x\<close> |
|
2290 |
using that by (simp add: divide_int_def sgn_if) |
|
2291 |
||
2292 |
lemma nat_div_distrib': |
|
2293 |
\<open>nat (x div y) = nat x div nat y\<close> if \<open>0 \<le> y\<close> |
|
2294 |
using that by (simp add: divide_int_def sgn_if) |
|
2295 |
||
2296 |
lemma nat_mod_distrib: \<comment> \<open>Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\<close> |
|
2297 |
\<open>nat (x mod y) = nat x mod nat y\<close> if \<open>0 \<le> x\<close> \<open>0 \<le> y\<close> |
|
2298 |
using that by (simp add: modulo_int_def sgn_if) |
|
2299 |
||
71157 | 2300 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2301 |
subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2302 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2303 |
lemma zdiv_mono1: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2304 |
\<open>a div b \<le> a' div b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2305 |
if \<open>a \<le> a'\<close> \<open>0 < b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2306 |
for a b b' :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2307 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2308 |
from \<open>a \<le> a'\<close> have \<open>b * (a div b) + a mod b \<le> b * (a' div b) + a' mod b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2309 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2310 |
then have \<open>b * (a div b) \<le> (a' mod b - a mod b) + b * (a' div b)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2311 |
by (simp add: algebra_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2312 |
moreover have \<open>a' mod b < b + a mod b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2313 |
by (rule less_le_trans [of _ b]) (use \<open>0 < b\<close> in simp_all) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2314 |
ultimately have \<open>b * (a div b) < b * (1 + a' div b)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2315 |
by (simp add: distrib_left) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2316 |
with \<open>0 < b\<close> have \<open>a div b < 1 + a' div b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2317 |
by (simp add: mult_less_cancel_left) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2318 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2319 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2320 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2321 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2322 |
lemma zdiv_mono1_neg: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2323 |
\<open>a' div b \<le> a div b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2324 |
if \<open>a \<le> a'\<close> \<open>b < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2325 |
for a a' b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2326 |
using that zdiv_mono1 [of \<open>- a'\<close> \<open>- a\<close> \<open>- b\<close>] by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2327 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2328 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2329 |
subsubsection \<open>Monotonicity in the Second Argument (Divisor)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2330 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2331 |
lemma zdiv_mono2: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2332 |
\<open>a div b \<le> a div b'\<close> if \<open>0 \<le> a\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> for a b b' :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2333 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2334 |
define q q' r r' where **: \<open>q = a div b\<close> \<open>q' = a div b'\<close> \<open>r = a mod b\<close> \<open>r' = a mod b'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2335 |
then have *: \<open>b * q + r = b' * q' + r'\<close> \<open>0 \<le> b' * q' + r'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2336 |
\<open>r' < b'\<close> \<open>0 \<le> r\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2337 |
using that by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2338 |
have \<open>0 < b' * (q' + 1)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2339 |
using * by (simp add: distrib_left) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2340 |
with * have \<open>0 \<le> q'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2341 |
by (simp add: zero_less_mult_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2342 |
moreover have \<open>b * q = r' - r + b' * q'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2343 |
using * by linarith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2344 |
ultimately have \<open>b * q < b * (q' + 1)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2345 |
using mult_right_mono * unfolding distrib_left by fastforce |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2346 |
with * have \<open>q \<le> q'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2347 |
by (simp add: mult_less_cancel_left_pos) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2348 |
with ** show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2349 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2350 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2351 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2352 |
lemma zdiv_mono2_neg: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2353 |
\<open>a div b' \<le> a div b\<close> if \<open>a < 0\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> for a b b' :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2354 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2355 |
define q q' r r' where **: \<open>q = a div b\<close> \<open>q' = a div b'\<close> \<open>r = a mod b\<close> \<open>r' = a mod b'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2356 |
then have *: \<open>b * q + r = b' * q' + r'\<close> \<open>b' * q' + r' < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2357 |
\<open>r < b\<close> \<open>0 \<le> r'\<close> \<open>0 < b'\<close> \<open>b' \<le> b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2358 |
using that by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2359 |
have \<open>b' * q' < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2360 |
using * by linarith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2361 |
with * have \<open>q' \<le> 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2362 |
by (simp add: mult_less_0_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2363 |
have \<open>b * q' \<le> b' * q'\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2364 |
by (simp add: \<open>q' \<le> 0\<close> * mult_right_mono_neg) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2365 |
then have "b * q' < b * (q + 1)" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2366 |
using * by (simp add: distrib_left) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2367 |
then have \<open>q' \<le> q\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2368 |
using * by (simp add: mult_less_cancel_left) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2369 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2370 |
by (simp add: **) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2371 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2372 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2373 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2374 |
subsubsection \<open>Quotients of Signs\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2375 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2376 |
lemma div_eq_minus1: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2377 |
\<open>0 < b \<Longrightarrow> - 1 div b = - 1\<close> for b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2378 |
by (simp add: divide_int_def) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2379 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2380 |
lemma zmod_minus1: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2381 |
\<open>0 < b \<Longrightarrow> - 1 mod b = b - 1\<close> for b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2382 |
by (auto simp add: modulo_int_def) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2383 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2384 |
lemma minus_mod_int_eq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2385 |
\<open>- k mod l = l - 1 - (k - 1) mod l\<close> if \<open>l \<ge> 0\<close> for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2386 |
proof (cases \<open>l = 0\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2387 |
case True |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2388 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2389 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2390 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2391 |
case False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2392 |
with that have \<open>l > 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2393 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2394 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2395 |
proof (cases \<open>l dvd k\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2396 |
case True |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2397 |
then obtain j where \<open>k = l * j\<close> .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2398 |
moreover have \<open>(l * j mod l - 1) mod l = l - 1\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2399 |
using \<open>l > 0\<close> by (simp add: zmod_minus1) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2400 |
then have \<open>(l * j - 1) mod l = l - 1\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2401 |
by (simp only: mod_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2402 |
ultimately show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2403 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2404 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2405 |
case False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2406 |
moreover have 1: \<open>0 < k mod l\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2407 |
using \<open>0 < l\<close> False le_less by fastforce |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2408 |
moreover have 2: \<open>k mod l < 1 + l\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2409 |
using \<open>0 < l\<close> pos_mod_bound[of l k] by linarith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2410 |
from 1 2 \<open>l > 0\<close> have \<open>(k mod l - 1) mod l = k mod l - 1\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2411 |
by (simp add: zmod_trivial_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2412 |
ultimately show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2413 |
by (simp only: zmod_zminus1_eq_if) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2414 |
(simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2415 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2416 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2417 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2418 |
lemma div_neg_pos_less0: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2419 |
\<open>a div b < 0\<close> if \<open>a < 0\<close> \<open>0 < b\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2420 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2421 |
have "a div b \<le> - 1 div b" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2422 |
using zdiv_mono1 that by auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2423 |
also have "... \<le> -1" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2424 |
by (simp add: that(2) div_eq_minus1) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2425 |
finally show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2426 |
by force |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2427 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2428 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2429 |
lemma div_nonneg_neg_le0: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2430 |
\<open>a div b \<le> 0\<close> if \<open>0 \<le> a\<close> \<open>b < 0\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2431 |
using that by (auto dest: zdiv_mono1_neg) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2432 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2433 |
lemma div_nonpos_pos_le0: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2434 |
\<open>a div b \<le> 0\<close> if \<open>a \<le> 0\<close> \<open>0 < b\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2435 |
using that by (auto dest: zdiv_mono1) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2436 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2437 |
text\<open>Now for some equivalences of the form \<open>a div b >=< 0 \<longleftrightarrow> \<dots>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2438 |
conditional upon the sign of \<open>a\<close> or \<open>b\<close>. There are many more. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2439 |
They should all be simp rules unless that causes too much search.\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2440 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2441 |
lemma pos_imp_zdiv_nonneg_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2442 |
\<open>0 \<le> a div b \<longleftrightarrow> 0 \<le> a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2443 |
if \<open>0 < b\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2444 |
proof |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2445 |
assume \<open>0 \<le> a div b\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2446 |
show \<open>0 \<le> a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2447 |
proof (rule ccontr) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2448 |
assume \<open>\<not> 0 \<le> a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2449 |
then have \<open>a < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2450 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2451 |
then have \<open>a div b < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2452 |
using that by (rule div_neg_pos_less0) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2453 |
with \<open>0 \<le> a div b\<close> show False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2454 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2455 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2456 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2457 |
assume "0 \<le> a" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2458 |
then have "0 div b \<le> a div b" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2459 |
using zdiv_mono1 that by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2460 |
then show "0 \<le> a div b" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2461 |
by auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2462 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2463 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2464 |
lemma neg_imp_zdiv_nonneg_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2465 |
\<open>0 \<le> a div b \<longleftrightarrow> a \<le> 0\<close> if \<open>b < 0\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2466 |
using that pos_imp_zdiv_nonneg_iff [of \<open>- b\<close> \<open>- a\<close>] by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2467 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2468 |
lemma pos_imp_zdiv_pos_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2469 |
\<open>0 < (i::int) div k \<longleftrightarrow> k \<le> i\<close> if \<open>0 < k\<close> for i k :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2470 |
using that pos_imp_zdiv_nonneg_iff [of k i] zdiv_eq_0_iff [of i k] by arith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2471 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2472 |
lemma pos_imp_zdiv_neg_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2473 |
\<open>a div b < 0 \<longleftrightarrow> a < 0\<close> if \<open>0 < b\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2474 |
\<comment> \<open>But not \<^prop>\<open>a div b \<le> 0 \<longleftrightarrow> a \<le> 0\<close>; consider \<^prop>\<open>a = 1\<close>, \<^prop>\<open>b = 2\<close> when \<^prop>\<open>a div b = 0\<close>.\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2475 |
using that by (simp add: pos_imp_zdiv_nonneg_iff flip: linorder_not_le) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2476 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2477 |
lemma neg_imp_zdiv_neg_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2478 |
\<comment> \<open>But not \<^prop>\<open>a div b \<le> 0 \<longleftrightarrow> 0 \<le> a\<close>; consider \<^prop>\<open>a = - 1\<close>, \<^prop>\<open>b = - 2\<close> when \<^prop>\<open>a div b = 0\<close>.\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2479 |
\<open>a div b < 0 \<longleftrightarrow> 0 < a\<close> if \<open>b < 0\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2480 |
using that by (simp add: neg_imp_zdiv_nonneg_iff flip: linorder_not_le) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2481 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2482 |
lemma nonneg1_imp_zdiv_pos_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2483 |
\<open>a div b > 0 \<longleftrightarrow> a \<ge> b \<and> b > 0\<close> if \<open>0 \<le> a\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2484 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2485 |
have "0 < a div b \<Longrightarrow> b \<le> a" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2486 |
using div_pos_pos_trivial[of a b] that by arith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2487 |
moreover have "0 < a div b \<Longrightarrow> b > 0" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2488 |
using that div_nonneg_neg_le0[of a b] by (cases "b=0"; force) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2489 |
moreover have "b \<le> a \<and> 0 < b \<Longrightarrow> 0 < a div b" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2490 |
using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2491 |
ultimately show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2492 |
by blast |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2493 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2494 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2495 |
lemma zmod_le_nonneg_dividend: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2496 |
\<open>m mod k \<le> m\<close> if \<open>(m::int) \<ge> 0\<close> for m k :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2497 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2498 |
from that have \<open>m > 0 \<or> m = 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2499 |
by auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2500 |
then show ?thesis proof |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2501 |
assume \<open>m = 0\<close> then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2502 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2503 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2504 |
assume \<open>m > 0\<close> then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2505 |
proof (cases k \<open>0::int\<close> rule: linorder_cases) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2506 |
case less |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2507 |
moreover define l where \<open>l = - k\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2508 |
ultimately have \<open>l > 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2509 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2510 |
with \<open>m > 0\<close> have \<open>int (nat m mod nat l) \<le> m\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2511 |
by (simp flip: le_nat_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2512 |
then have \<open>int (nat m mod nat l) - l \<le> m\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2513 |
using \<open>l > 0\<close> by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2514 |
with \<open>m > 0\<close> \<open>l > 0\<close> show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2515 |
by (simp add: modulo_int_def l_def flip: le_nat_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2516 |
qed (simp_all add: modulo_int_def flip: le_nat_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2517 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2518 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2519 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2520 |
lemma sgn_div_eq_sgn_mult: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2521 |
\<open>sgn (k div l) = of_bool (k div l \<noteq> 0) * sgn (k * l)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2522 |
for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2523 |
proof (cases \<open>k div l = 0\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2524 |
case True |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2525 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2526 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2527 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2528 |
case False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2529 |
have \<open>0 \<le> \<bar>k\<bar> div \<bar>l\<bar>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2530 |
by (cases \<open>l = 0\<close>) (simp_all add: pos_imp_zdiv_nonneg_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2531 |
then have \<open>\<bar>k\<bar> div \<bar>l\<bar> \<noteq> 0 \<longleftrightarrow> 0 < \<bar>k\<bar> div \<bar>l\<bar>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2532 |
by (simp add: less_le) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2533 |
also have \<open>\<dots> \<longleftrightarrow> \<bar>k\<bar> \<ge> \<bar>l\<bar>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2534 |
using False nonneg1_imp_zdiv_pos_iff by auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2535 |
finally have *: \<open>\<bar>k\<bar> div \<bar>l\<bar> \<noteq> 0 \<longleftrightarrow> \<bar>l\<bar> \<le> \<bar>k\<bar>\<close> . |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2536 |
show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2537 |
using \<open>0 \<le> \<bar>k\<bar> div \<bar>l\<bar>\<close> False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2538 |
by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2539 |
sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2540 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2541 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2542 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2543 |
subsubsection \<open>Further properties\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2544 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2545 |
lemma div_int_pos_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2546 |
"k div l \<ge> 0 \<longleftrightarrow> k = 0 \<or> l = 0 \<or> k \<ge> 0 \<and> l \<ge> 0 |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2547 |
\<or> k < 0 \<and> l < 0" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2548 |
for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2549 |
proof (cases "k = 0 \<or> l = 0") |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2550 |
case False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2551 |
then have *: "k \<noteq> 0" "l \<noteq> 0" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2552 |
by auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2553 |
then have "0 \<le> k div l \<Longrightarrow> \<not> k < 0 \<Longrightarrow> 0 \<le> l" |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2554 |
by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2555 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2556 |
using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2557 |
qed auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2558 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2559 |
lemma mod_int_pos_iff: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2560 |
\<open>k mod l \<ge> 0 \<longleftrightarrow> l dvd k \<or> l = 0 \<and> k \<ge> 0 \<or> l > 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2561 |
for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2562 |
proof (cases "l > 0") |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2563 |
case False |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2564 |
then show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2565 |
by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \<open>auto simp add: le_less not_less\<close>) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2566 |
qed auto |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2567 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2568 |
lemma abs_div: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2569 |
\<open>\<bar>x div y\<bar> = \<bar>x\<bar> div \<bar>y\<bar>\<close> if \<open>y dvd x\<close> for x y :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2570 |
using that by (cases \<open>y = 0\<close>) (auto simp add: abs_mult) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2571 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2572 |
lemma int_power_div_base: \<^marker>\<open>contributor \<open>Matthias Daum\<close>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2573 |
\<open>k ^ m div k = k ^ (m - Suc 0)\<close> if \<open>0 < m\<close> \<open>0 < k\<close> for k :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2574 |
using that by (cases m) simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2575 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2576 |
lemma int_div_less_self: \<^marker>\<open>contributor \<open>Matthias Daum\<close>\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2577 |
\<open>x div k < x\<close> if \<open>0 < x\<close> \<open>1 < k\<close> for x k :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2578 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2579 |
from that have \<open>nat (x div k) = nat x div nat k\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2580 |
by (simp add: nat_div_distrib) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2581 |
also from that have \<open>nat x div nat k < nat x\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2582 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2583 |
finally show ?thesis |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2584 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2585 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2586 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2587 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2588 |
subsubsection \<open>Computing \<open>div\<close> and \<open>mod\<close> by shifting\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2589 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2590 |
lemma div_pos_geq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2591 |
\<open>k div l = (k - l) div l + 1\<close> if \<open>0 < l\<close> \<open>l \<le> k\<close> for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2592 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2593 |
have "k = (k - l) + l" by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2594 |
then obtain j where k: "k = j + l" .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2595 |
with that show ?thesis by (simp add: div_add_self2) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2596 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2597 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2598 |
lemma mod_pos_geq: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2599 |
\<open>k mod l = (k - l) mod l\<close> if \<open>0 < l\<close> \<open>l \<le> k\<close> for k l :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2600 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2601 |
have "k = (k - l) + l" by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2602 |
then obtain j where k: "k = j + l" .. |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2603 |
with that show ?thesis by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2604 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2605 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2606 |
lemma pos_zdiv_mult_2: \<open>(1 + 2 * b) div (2 * a) = b div a\<close> (is ?Q) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2607 |
and pos_zmod_mult_2: \<open>(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\<close> (is ?R) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2608 |
if \<open>0 \<le> a\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2609 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2610 |
have \<open>((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\<close> |
76245 | 2611 |
proof (induction rule: euclidean_relation_intI) |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2612 |
case by0 |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2613 |
then show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2614 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2615 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2616 |
case divides |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2617 |
have \<open>2 dvd (2 * a)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2618 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2619 |
then have \<open>2 dvd (1 + 2 * b)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2620 |
using \<open>2 * a dvd 1 + 2 * b\<close> by (rule dvd_trans) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2621 |
then have \<open>2 dvd (1 + b * 2)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2622 |
by (simp add: ac_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2623 |
then have \<open>is_unit (2 :: int)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2624 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2625 |
then show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2626 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2627 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2628 |
case euclidean_relation |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2629 |
with that have \<open>a > 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2630 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2631 |
moreover have \<open>b mod a < a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2632 |
using \<open>a > 0\<close> by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2633 |
then have \<open>1 + 2 * (b mod a) < 2 * a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2634 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2635 |
moreover have \<open>2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2636 |
by (simp only: algebra_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2637 |
moreover have \<open>0 \<le> 2 * (b mod a)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2638 |
using \<open>a > 0\<close> by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2639 |
ultimately show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2640 |
by (simp add: algebra_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2641 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2642 |
then show ?Q and ?R |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2643 |
by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2644 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2645 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2646 |
lemma neg_zdiv_mult_2: \<open>(1 + 2 * b) div (2 * a) = (b + 1) div a\<close> (is ?Q) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2647 |
and neg_zmod_mult_2: \<open>(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\<close> (is ?R) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2648 |
if \<open>a \<le> 0\<close> for a b :: int |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2649 |
proof - |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2650 |
have \<open>((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\<close> |
76245 | 2651 |
proof (induction rule: euclidean_relation_intI) |
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2652 |
case by0 |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2653 |
then show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2654 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2655 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2656 |
case divides |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2657 |
have \<open>2 dvd (2 * a)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2658 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2659 |
then have \<open>2 dvd (1 + 2 * b)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2660 |
using \<open>2 * a dvd 1 + 2 * b\<close> by (rule dvd_trans) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2661 |
then have \<open>2 dvd (1 + b * 2)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2662 |
by (simp add: ac_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2663 |
then have \<open>is_unit (2 :: int)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2664 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2665 |
then show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2666 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2667 |
next |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2668 |
case euclidean_relation |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2669 |
with that have \<open>a < 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2670 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2671 |
moreover have \<open>(b + 1) mod a > a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2672 |
using \<open>a < 0\<close> by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2673 |
then have \<open>2 * ((b + 1) mod a) > 1 + 2 * a\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2674 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2675 |
moreover have \<open>((1 + b) mod a) \<le> 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2676 |
using \<open>a < 0\<close> by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2677 |
then have \<open>2 * ((1 + b) mod a) \<le> 0\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2678 |
by simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2679 |
moreover have \<open>2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) = |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2680 |
2 * ((1 + b) div a * a + (1 + b) mod a)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2681 |
by (simp only: algebra_simps) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2682 |
ultimately show ?case |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2683 |
by (simp add: algebra_simps sgn_mult abs_mult) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2684 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2685 |
then show ?Q and ?R |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2686 |
by simp_all |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2687 |
qed |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2688 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2689 |
lemma zdiv_numeral_Bit0 [simp]: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2690 |
\<open>numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2691 |
numeral v div (numeral w :: int)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2692 |
unfolding numeral.simps unfolding mult_2 [symmetric] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2693 |
by (rule div_mult_mult1) simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2694 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2695 |
lemma zdiv_numeral_Bit1 [simp]: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2696 |
\<open>numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2697 |
(numeral v div (numeral w :: int))\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2698 |
unfolding numeral.simps |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2699 |
unfolding mult_2 [symmetric] add.commute [of _ 1] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2700 |
by (rule pos_zdiv_mult_2) simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2701 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2702 |
lemma zmod_numeral_Bit0 [simp]: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2703 |
\<open>numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2704 |
(2::int) * (numeral v mod numeral w)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2705 |
unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2706 |
unfolding mult_2 [symmetric] by (rule mod_mult_mult1) |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2707 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2708 |
lemma zmod_numeral_Bit1 [simp]: |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2709 |
\<open>numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2710 |
2 * (numeral v mod numeral w) + (1::int)\<close> |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2711 |
unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2712 |
unfolding mult_2 [symmetric] add.commute [of _ 1] |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2713 |
by (rule pos_zmod_mult_2) simp |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2714 |
|
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
2715 |
|
76387 | 2716 |
subsection \<open>Code generation\<close> |
75937 | 2717 |
|
2718 |
context |
|
2719 |
begin |
|
2720 |
||
2721 |
qualified definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" |
|
2722 |
where "divmod_nat m n = (m div n, m mod n)" |
|
2723 |
||
2724 |
qualified lemma divmod_nat_if [code]: |
|
2725 |
"divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else |
|
2726 |
let (q, r) = divmod_nat (m - n) n in (Suc q, r))" |
|
2727 |
by (simp add: divmod_nat_def prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) |
|
2728 |
||
2729 |
qualified lemma [code]: |
|
2730 |
"m div n = fst (divmod_nat m n)" |
|
2731 |
"m mod n = snd (divmod_nat m n)" |
|
2732 |
by (simp_all add: divmod_nat_def) |
|
2733 |
||
2734 |
end |
|
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
2735 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
2736 |
code_identifier |
77061
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents:
76387
diff
changeset
|
2737 |
code_module Euclidean_Rings \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
2738 |
|
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66807
diff
changeset
|
2739 |
end |