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