author | haftmann |
Thu, 29 Sep 2022 14:03:40 +0000 | |
changeset 76224 | 64e8d4afcf10 |
parent 76141 | e7497a1de8b9 |
child 76231 | 8a48e18f081e |
permissions | -rw-r--r-- |
3366 | 1 |
(* Title: HOL/Divides.thy |
18154 | 2 |
*) |
3366 | 3 |
|
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
4 |
section \<open>Lemmas of doubtful value\<close> |
3366 | 5 |
|
15131 | 6 |
theory Divides |
66817 | 7 |
imports Parity |
15131 | 8 |
begin |
3366 | 9 |
|
75936 | 10 |
class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom + |
11 |
assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0" |
|
12 |
and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a" |
|
13 |
and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0" |
|
14 |
and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a" |
|
15 |
and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b" |
|
16 |
and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b" |
|
17 |
and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b" |
|
18 |
and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c" |
|
19 |
assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b" |
|
20 |
begin |
|
21 |
||
22 |
lemma divmod_digit_1: |
|
23 |
assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)" |
|
24 |
shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") |
|
25 |
and "a mod (2 * b) - b = a mod b" (is "?Q") |
|
26 |
proof - |
|
27 |
from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a" |
|
28 |
by (auto intro: trans) |
|
29 |
with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive) |
|
30 |
then have [simp]: "1 \<le> a div b" by (simp add: discrete) |
|
31 |
with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound) |
|
32 |
define w where "w = a div b mod 2" |
|
33 |
then have w_exhaust: "w = 0 \<or> w = 1" by auto |
|
34 |
have mod_w: "a mod (2 * b) = a mod b + b * w" |
|
35 |
by (simp add: w_def mod_mult2_eq ac_simps) |
|
36 |
from assms w_exhaust have "w = 1" |
|
37 |
using mod_less by (auto simp add: mod_w) |
|
38 |
with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp |
|
39 |
have "2 * (a div (2 * b)) = a div b - w" |
|
40 |
by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) |
|
41 |
with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp |
|
42 |
then show ?P and ?Q |
|
43 |
by (simp_all add: div mod add_implies_diff [symmetric]) |
|
44 |
qed |
|
45 |
||
46 |
lemma divmod_digit_0: |
|
47 |
assumes "0 < b" and "a mod (2 * b) < b" |
|
48 |
shows "2 * (a div (2 * b)) = a div b" (is "?P") |
|
49 |
and "a mod (2 * b) = a mod b" (is "?Q") |
|
50 |
proof - |
|
51 |
define w where "w = a div b mod 2" |
|
52 |
then have w_exhaust: "w = 0 \<or> w = 1" by auto |
|
53 |
have mod_w: "a mod (2 * b) = a mod b + b * w" |
|
54 |
by (simp add: w_def mod_mult2_eq ac_simps) |
|
55 |
moreover have "b \<le> a mod b + b" |
|
56 |
proof - |
|
57 |
from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast |
|
58 |
then have "0 + b \<le> a mod b + b" by (rule add_right_mono) |
|
59 |
then show ?thesis by simp |
|
60 |
qed |
|
61 |
moreover note assms w_exhaust |
|
62 |
ultimately have "w = 0" by auto |
|
63 |
with mod_w have mod: "a mod (2 * b) = a mod b" by simp |
|
64 |
have "2 * (a div (2 * b)) = a div b - w" |
|
65 |
by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) |
|
66 |
with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp |
|
67 |
then show ?P and ?Q |
|
68 |
by (simp_all add: div mod) |
|
69 |
qed |
|
70 |
||
71 |
lemma mod_double_modulus: |
|
72 |
assumes "m > 0" "x \<ge> 0" |
|
73 |
shows "x mod (2 * m) = x mod m \<or> x mod (2 * m) = x mod m + m" |
|
74 |
proof (cases "x mod (2 * m) < m") |
|
75 |
case True |
|
76 |
thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto |
|
77 |
next |
|
78 |
case False |
|
79 |
hence *: "x mod (2 * m) - m = x mod m" |
|
80 |
using assms by (intro divmod_digit_1) auto |
|
81 |
hence "x mod (2 * m) = x mod m + m" |
|
82 |
by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto) |
|
83 |
thus ?thesis by simp |
|
84 |
qed |
|
85 |
||
86 |
end |
|
87 |
||
88 |
hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq |
|
89 |
||
90 |
instance nat :: unique_euclidean_semiring_numeral |
|
91 |
by standard |
|
92 |
(auto simp add: div_greater_zero_iff div_mult2_eq mod_mult2_eq) |
|
93 |
||
94 |
instance int :: unique_euclidean_semiring_numeral |
|
95 |
by standard (auto intro: zmod_le_nonneg_dividend simp add: |
|
96 |
pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) |
|
97 |
||
68631 | 98 |
lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66806
diff
changeset
|
99 |
by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66806
diff
changeset
|
100 |
|
68631 | 101 |
lemma mod_geq: "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66806
diff
changeset
|
102 |
by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>) |
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66806
diff
changeset
|
103 |
|
68631 | 104 |
lemma mod_eq_0D: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: nat |
105 |
using that by (auto simp add: mod_eq_0_iff_dvd) |
|
66816
212a3334e7da
more fundamental definition of div and mod on int
haftmann
parents:
66815
diff
changeset
|
106 |
|
69695 | 107 |
lemma pos_mod_conj: "0 < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b" for a b :: int |
108 |
by simp |
|
109 |
||
110 |
lemma neg_mod_conj: "b < 0 \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b" for a b :: int |
|
111 |
by simp |
|
112 |
||
113 |
lemma zmod_eq_0_iff: "m mod d = 0 \<longleftrightarrow> (\<exists>q. m = d * q)" for m d :: int |
|
114 |
by (auto simp add: mod_eq_0_iff_dvd) |
|
115 |
||
116 |
(* REVISIT: should this be generalized to all semiring_div types? *) |
|
117 |
lemma zmod_eq_0D [dest!]: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: int |
|
118 |
using that by auto |
|
119 |
||
75937 | 120 |
lemma div_positive_int: |
121 |
"k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int |
|
122 |
using that by (simp add: nonneg1_imp_zdiv_pos_iff) |
|
123 |
||
76224
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
124 |
code_identifier |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
125 |
code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith |
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
haftmann
parents:
76141
diff
changeset
|
126 |
|
33361
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
haftmann
parents:
33340
diff
changeset
|
127 |
end |