| author | Fabian Huch <huch@in.tum.de> | 
| Thu, 27 Jun 2024 09:41:16 +0200 | |
| changeset 80414 | 4b10ae56ed01 | 
| parent 78937 | 5e6b195eee83 | 
| child 80452 | 0303b3a0edde | 
| 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 | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78935diff
changeset | 10 | class unique_euclidean_semiring_numeral = linordered_euclidean_semiring + discrete_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 | ||
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
78669diff
changeset | 20 | hide_fact (open) div_less mod_less div_positive mod_less_eq_dividend pos_mod_bound pos_mod_sign mod_mult2_eq div_mult2_eq | 
| 76231 | 21 | |
| 22 | context unique_euclidean_semiring_numeral | |
| 75936 | 23 | begin | 
| 24 | ||
| 78669 | 25 | context | 
| 26 | begin | |
| 27 | ||
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
78669diff
changeset | 28 | qualified lemma discrete [no_atp]: | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
78669diff
changeset | 29 | "a < b \<longleftrightarrow> a + 1 \<le> b" | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
78669diff
changeset | 30 | by (fact less_iff_succ_less_eq) | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
78669diff
changeset | 31 | |
| 78669 | 32 | qualified lemma divmod_digit_1 [no_atp]: | 
| 75936 | 33 | assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)" | 
| 34 | shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") | |
| 35 | and "a mod (2 * b) - b = a mod b" (is "?Q") | |
| 36 | proof - | |
| 37 | from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a" | |
| 38 | by (auto intro: trans) | |
| 39 | with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive) | |
| 40 | then have [simp]: "1 \<le> a div b" by (simp add: discrete) | |
| 41 | with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound) | |
| 42 | define w where "w = a div b mod 2" | |
| 43 | then have w_exhaust: "w = 0 \<or> w = 1" by auto | |
| 44 | have mod_w: "a mod (2 * b) = a mod b + b * w" | |
| 45 | by (simp add: w_def mod_mult2_eq ac_simps) | |
| 46 | from assms w_exhaust have "w = 1" | |
| 47 | using mod_less by (auto simp add: mod_w) | |
| 48 | with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp | |
| 49 | have "2 * (a div (2 * b)) = a div b - w" | |
| 50 | by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) | |
| 51 | with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp | |
| 52 | then show ?P and ?Q | |
| 53 | by (simp_all add: div mod add_implies_diff [symmetric]) | |
| 54 | qed | |
| 55 | ||
| 78669 | 56 | qualified lemma divmod_digit_0 [no_atp]: | 
| 75936 | 57 | assumes "0 < b" and "a mod (2 * b) < b" | 
| 58 | shows "2 * (a div (2 * b)) = a div b" (is "?P") | |
| 59 | and "a mod (2 * b) = a mod b" (is "?Q") | |
| 60 | proof - | |
| 61 | define w where "w = a div b mod 2" | |
| 62 | then have w_exhaust: "w = 0 \<or> w = 1" by auto | |
| 63 | have mod_w: "a mod (2 * b) = a mod b + b * w" | |
| 64 | by (simp add: w_def mod_mult2_eq ac_simps) | |
| 65 | moreover have "b \<le> a mod b + b" | |
| 66 | proof - | |
| 67 | from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast | |
| 68 | then have "0 + b \<le> a mod b + b" by (rule add_right_mono) | |
| 69 | then show ?thesis by simp | |
| 70 | qed | |
| 71 | moreover note assms w_exhaust | |
| 72 | ultimately have "w = 0" by auto | |
| 73 | with mod_w have mod: "a mod (2 * b) = a mod b" by simp | |
| 74 | have "2 * (a div (2 * b)) = a div b - w" | |
| 75 | by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) | |
| 76 | with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp | |
| 77 | then show ?P and ?Q | |
| 78 | by (simp_all add: div mod) | |
| 79 | qed | |
| 80 | ||
| 78669 | 81 | qualified lemma mod_double_modulus [no_atp]: | 
| 75936 | 82 | assumes "m > 0" "x \<ge> 0" | 
| 83 | shows "x mod (2 * m) = x mod m \<or> x mod (2 * m) = x mod m + m" | |
| 84 | proof (cases "x mod (2 * m) < m") | |
| 85 | case True | |
| 86 | thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto | |
| 87 | next | |
| 88 | case False | |
| 89 | hence *: "x mod (2 * m) - m = x mod m" | |
| 90 | using assms by (intro divmod_digit_1) auto | |
| 91 | hence "x mod (2 * m) = x mod m + m" | |
| 92 | by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto) | |
| 93 | thus ?thesis by simp | |
| 94 | qed | |
| 95 | ||
| 96 | end | |
| 97 | ||
| 78669 | 98 | end | 
| 99 | ||
| 75936 | 100 | instance nat :: unique_euclidean_semiring_numeral | 
| 101 | by standard | |
| 102 | (auto simp add: div_greater_zero_iff div_mult2_eq mod_mult2_eq) | |
| 103 | ||
| 104 | instance int :: unique_euclidean_semiring_numeral | |
| 105 | by standard (auto intro: zmod_le_nonneg_dividend simp add: | |
| 106 | pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) | |
| 107 | ||
| 78669 | 108 | context | 
| 109 | begin | |
| 110 | ||
| 69695 | 111 | (* REVISIT: should this be generalized to all semiring_div types? *) | 
| 78669 | 112 | qualified lemma zmod_eq_0D [dest!]: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: int | 
| 69695 | 113 | using that by auto | 
| 114 | ||
| 78669 | 115 | qualified lemma div_geq [no_atp]: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat | 
| 76231 | 116 | by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>) | 
| 117 | ||
| 78669 | 118 | qualified lemma mod_geq [no_atp]: "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat | 
| 76231 | 119 | by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>) | 
| 120 | ||
| 78669 | 121 | qualified lemma mod_eq_0D [no_atp]: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: nat | 
| 76231 | 122 | using that by (auto simp add: mod_eq_0_iff_dvd) | 
| 123 | ||
| 78669 | 124 | qualified lemma pos_mod_conj [no_atp]: "0 < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b" for a b :: int | 
| 76231 | 125 | by simp | 
| 126 | ||
| 78669 | 127 | qualified lemma neg_mod_conj [no_atp]: "b < 0 \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b" for a b :: int | 
| 76231 | 128 | by simp | 
| 129 | ||
| 78669 | 130 | qualified lemma zmod_eq_0_iff [no_atp]: "m mod d = 0 \<longleftrightarrow> (\<exists>q. m = d * q)" for m d :: int | 
| 76231 | 131 | by (auto simp add: mod_eq_0_iff_dvd) | 
| 132 | ||
| 78669 | 133 | qualified lemma div_positive_int [no_atp]: | 
| 75937 | 134 | "k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int | 
| 135 | using that by (simp add: nonneg1_imp_zdiv_pos_iff) | |
| 136 | ||
| 78669 | 137 | end | 
| 138 | ||
| 76224 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
76141diff
changeset | 139 | code_identifier | 
| 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
76141diff
changeset | 140 | 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 | 141 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 142 | end |