diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,127 +1,134 @@ (* Title: HOL/Divides.thy *) -section \Lemmas of doubtful value\ +section \Misc lemmas on division, to be sorted out finally\ theory Divides imports Parity begin class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom + - assumes div_less: "0 \ a \ a < b \ a div b = 0" - and mod_less: " 0 \ a \ a < b \ a mod b = a" - and div_positive: "0 < b \ b \ a \ a div b > 0" - and mod_less_eq_dividend: "0 \ a \ a mod b \ a" - and pos_mod_bound: "0 < b \ a mod b < b" - and pos_mod_sign: "0 < b \ 0 \ a mod b" - and mod_mult2_eq: "0 \ c \ a mod (b * c) = b * (a div b mod c) + a mod b" - and div_mult2_eq: "0 \ c \ a div (b * c) = a div b div c" - assumes discrete: "a < b \ a + 1 \ b" + assumes div_less [no_atp]: "0 \ a \ a < b \ a div b = 0" + and mod_less [no_atp]: " 0 \ a \ a < b \ a mod b = a" + and div_positive [no_atp]: "0 < b \ b \ a \ a div b > 0" + and mod_less_eq_dividend [no_atp]: "0 \ a \ a mod b \ a" + and pos_mod_bound [no_atp]: "0 < b \ a mod b < b" + and pos_mod_sign [no_atp]: "0 < b \ 0 \ a mod b" + and mod_mult2_eq [no_atp]: "0 \ c \ a mod (b * c) = b * (a div b mod c) + a mod b" + and div_mult2_eq [no_atp]: "0 \ c \ a div (b * c) = a div b div c" + assumes discrete [no_atp]: "a < b \ a + 1 \ b" + +hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq + +context unique_euclidean_semiring_numeral begin -lemma divmod_digit_1: +context +begin + +lemma divmod_digit_1 [no_atp]: assumes "0 \ a" "0 < b" and "b \ a mod (2 * b)" shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") and "a mod (2 * b) - b = a mod b" (is "?Q") proof - from assms mod_less_eq_dividend [of a "2 * b"] have "b \ a" by (auto intro: trans) with \0 < b\ have "0 < a div b" by (auto intro: div_positive) then have [simp]: "1 \ a div b" by (simp add: discrete) with \0 < b\ have mod_less: "a mod b < b" by (simp add: pos_mod_bound) define w where "w = a div b mod 2" then have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) from assms w_exhaust have "w = 1" using mod_less by (auto simp add: mod_w) with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp have "2 * (a div (2 * b)) = a div b - w" by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) with \w = 1\ have div: "2 * (a div (2 * b)) = a div b - 1" by simp then show ?P and ?Q by (simp_all add: div mod add_implies_diff [symmetric]) qed -lemma divmod_digit_0: +lemma divmod_digit_0 [no_atp]: assumes "0 < b" and "a mod (2 * b) < b" shows "2 * (a div (2 * b)) = a div b" (is "?P") and "a mod (2 * b) = a mod b" (is "?Q") proof - define w where "w = a div b mod 2" then have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) moreover have "b \ a mod b + b" proof - from \0 < b\ pos_mod_sign have "0 \ a mod b" by blast then have "0 + b \ a mod b + b" by (rule add_right_mono) then show ?thesis by simp qed moreover note assms w_exhaust ultimately have "w = 0" by auto with mod_w have mod: "a mod (2 * b) = a mod b" by simp have "2 * (a div (2 * b)) = a div b - w" by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) with \w = 0\ have div: "2 * (a div (2 * b)) = a div b" by simp then show ?P and ?Q by (simp_all add: div mod) qed -lemma mod_double_modulus: +lemma mod_double_modulus: \\This is actually useful and should be transferred to a suitable type class\ assumes "m > 0" "x \ 0" shows "x mod (2 * m) = x mod m \ x mod (2 * m) = x mod m + m" proof (cases "x mod (2 * m) < m") case True thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto next case False hence *: "x mod (2 * m) - m = x mod m" using assms by (intro divmod_digit_1) auto hence "x mod (2 * m) = x mod m + m" by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto) thus ?thesis by simp qed end -hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq +end instance nat :: unique_euclidean_semiring_numeral by standard (auto simp add: div_greater_zero_iff div_mult2_eq mod_mult2_eq) instance int :: unique_euclidean_semiring_numeral by standard (auto intro: zmod_le_nonneg_dividend simp add: pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) -lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \ m < n" for m n :: nat - by (rule le_div_geq) (use that in \simp_all add: not_less\) - -lemma mod_geq: "m mod n = (m - n) mod n" if "\ m < n" for m n :: nat - by (rule le_mod_geq) (use that in \simp add: not_less\) - -lemma mod_eq_0D: "\q. m = d * q" if "m mod d = 0" for m d :: nat - using that by (auto simp add: mod_eq_0_iff_dvd) - -lemma pos_mod_conj: "0 < b \ 0 \ a mod b \ a mod b < b" for a b :: int - by simp - -lemma neg_mod_conj: "b < 0 \ a mod b \ 0 \ b < a mod b" for a b :: int - by simp - -lemma zmod_eq_0_iff: "m mod d = 0 \ (\q. m = d * q)" for m d :: int - by (auto simp add: mod_eq_0_iff_dvd) - (* REVISIT: should this be generalized to all semiring_div types? *) lemma zmod_eq_0D [dest!]: "\q. m = d * q" if "m mod d = 0" for m d :: int using that by auto -lemma div_positive_int: +lemma div_geq [no_atp]: "m div n = Suc ((m - n) div n)" if "0 < n" and " \ m < n" for m n :: nat + by (rule le_div_geq) (use that in \simp_all add: not_less\) + +lemma mod_geq [no_atp]: "m mod n = (m - n) mod n" if "\ m < n" for m n :: nat + by (rule le_mod_geq) (use that in \simp add: not_less\) + +lemma mod_eq_0D [no_atp]: "\q. m = d * q" if "m mod d = 0" for m d :: nat + using that by (auto simp add: mod_eq_0_iff_dvd) + +lemma pos_mod_conj [no_atp]: "0 < b \ 0 \ a mod b \ a mod b < b" for a b :: int + by simp + +lemma neg_mod_conj [no_atp]: "b < 0 \ a mod b \ 0 \ b < a mod b" for a b :: int + by simp + +lemma zmod_eq_0_iff [no_atp]: "m mod d = 0 \ (\q. m = d * q)" for m d :: int + by (auto simp add: mod_eq_0_iff_dvd) + +lemma div_positive_int [no_atp]: "k div l > 0" if "k \ l" and "l > 0" for k l :: int using that by (simp add: nonneg1_imp_zdiv_pos_iff) code_identifier code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,3319 +1,3376 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (use assms in simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule iffI; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") auto lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ lemma power_mod [mod_simps]: "((a mod b) ^ n) mod b = (a ^ n) mod b" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" by (simp add: mod_mult_right_eq) with Suc show ?case by (simp add: mod_mult_left_eq mod_mult_right_eq) qed lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp qed end class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel begin subclass idom_divide .. lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" using div_mult_mult1 [of "- 1" a b] by simp lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" using mod_mult_mult1 [of "- 1" a b] by simp lemma div_minus_right: "a div (- b) = (- a) div b" using div_minus_minus [of "- a" b] by simp lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" using mod_minus_minus [of "- a" b] by simp lemma div_minus1_right [simp]: "a div (- 1) = - a" using div_minus_right [of a 1] by simp lemma mod_minus1_right [simp]: "a mod (- 1) = 0" using mod_minus_right [of a 1] by simp text \Negation respects modular equivalence.\ lemma mod_minus_eq [mod_simps]: "(- (a mod b)) mod b = (- a) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: div_mult_mod_eq) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. then show ?thesis by (simp add: mod_minus_eq) qed text \Subtraction respects modular equivalence.\ lemma mod_diff_left_eq [mod_simps]: "(a mod c - b) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp lemma mod_diff_right_eq [mod_simps]: "(a - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_eq: "(a mod c - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp lemma minus_mod_self2 [simp]: "(a - b) mod b = a mod b" using mod_diff_right_eq [of a b b] by (simp add: mod_diff_right_eq) lemma minus_mod_self1 [simp]: "(b - a) mod b = - a mod b" using mod_add_self2 [of "- a" b] by simp lemma mod_eq_dvd_iff: "a mod c = b mod c \ c dvd a - b" (is "?P \ ?Q") proof assume ?P then have "(a mod c - b mod c) mod c = 0" by simp then show ?Q by (simp add: dvd_eq_mod_eq_0 mod_simps) next assume ?Q then obtain d where d: "a - b = c * d" .. then have "a = c * d + b" by (simp add: algebra_simps) then show ?P by simp qed lemma mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: \euclidean_size (a * b) = euclidean_size a * euclidean_size b\ fixes division_segment :: \'a \ 'a\ assumes is_unit_division_segment [simp]: \is_unit (division_segment a)\ and division_segment_mult: \a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b\ and division_segment_mod: \b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b\ assumes div_bounded: \b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q\ begin lemma division_segment_not_0 [simp]: \division_segment a \ 0\ using is_unit_division_segment [of a] is_unitE [of \division_segment a\] by blast lemma euclidean_relationI [case_names by0 divides euclidean_relation]: \(a div b, a mod b) = (q, r)\ if by0: \b = 0 \ q = 0 \ r = a\ and divides: \b \ 0 \ b dvd a \ r = 0 \ a = q * b\ and euclidean_relation: \b \ 0 \ \ b dvd a \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ a = q * b + r\ proof (cases \b = 0\) case True with by0 show ?thesis by simp next case False show ?thesis proof (cases \b dvd a\) case True with \b \ 0\ divides show ?thesis by simp next case False with \b \ 0\ euclidean_relation have \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ \a = q * b + r\ by simp_all from \b \ 0\ \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ have \(q * b + r) div b = q\ by (rule div_bounded) with \a = q * b + r\ have \q = a div b\ by simp from \a = q * b + r\ have \a div b * b + a mod b = q * b + r\ by (simp add: div_mult_mod_eq) with \q = a div b\ have \q * b + a mod b = q * b + r\ by simp then have \r = a mod b\ by simp with \q = a div b\ show ?thesis by simp qed qed subclass euclidean_semiring_cancel proof fix a b c assume \b \ 0\ have \((a + c * b) div b, (a + c * b) mod b) = (c + a div b, a mod b)\ proof (cases b \c + a div b\ \a mod b\ \a + c * b\ rule: euclidean_relationI) case by0 with \b \ 0\ show ?case by simp next case divides then show ?case by (simp add: algebra_simps dvd_add_left_iff) next case euclidean_relation then have \\ b dvd a\ by (simp add: dvd_add_left_iff) have \a mod b + (b * c + b * (a div b)) = b * c + ((a div b) * b + a mod b)\ by (simp add: ac_simps) with \b \ 0\ have *: \a mod b + (b * c + b * (a div b)) = b * c + a\ by (simp add: div_mult_mod_eq) from \\ b dvd a\ euclidean_relation show ?case by (simp_all add: algebra_simps division_segment_mod mod_size_less *) qed then show \(a + c * b) div b = c + a div b\ by simp next fix a b c assume \c \ 0\ have \((c * a) div (c * b), (c * a) mod (c * b)) = (a div b, c * (a mod b))\ proof (cases \c * b\ \a div b\ \c * (a mod b)\ \c * a\ rule: euclidean_relationI) case by0 with \c \ 0\ show ?case by simp next case divides then show ?case by (auto simp add: algebra_simps) next case euclidean_relation then have \b \ 0\ \a mod b \ 0\ by (simp_all add: mod_eq_0_iff_dvd) have \c * (a mod b) + b * (c * (a div b)) = c * ((a div b) * b + a mod b)\ by (simp add: algebra_simps) with \b \ 0\ have *: \c * (a mod b) + b * (c * (a div b)) = c * a\ by (simp add: div_mult_mod_eq) from \b \ 0\ \c \ 0\ have \euclidean_size c * euclidean_size (a mod b) < euclidean_size c * euclidean_size b\ using mod_size_less [of b a] by simp with euclidean_relation \b \ 0\ \a mod b \ 0\ show ?case by (simp add: algebra_simps division_segment_mult division_segment_mod euclidean_size_mult *) qed then show \(c * a) div (c * b) = a div b\ by simp qed lemma div_eq_0_iff: \a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0\ (is "_ \ ?P") if \division_segment a = division_segment b\ proof (cases \a = 0 \ b = 0\) case True then show ?thesis by auto next case False then have \a \ 0\ \b \ 0\ by simp_all have \a div b = 0 \ euclidean_size a < euclidean_size b\ proof assume \a div b = 0\ then have \a mod b = a\ using div_mult_mod_eq [of a b] by simp with \b \ 0\ mod_size_less [of b a] show \euclidean_size a < euclidean_size b\ by simp next assume \euclidean_size a < euclidean_size b\ have \(a div b, a mod b) = (0, a)\ proof (cases b 0 a a rule: euclidean_relationI) case by0 show ?case by simp next case divides with \euclidean_size a < euclidean_size b\ show ?case using dvd_imp_size_le [of b a] \a \ 0\ by simp next case euclidean_relation with \euclidean_size a < euclidean_size b\ that show ?case by simp qed then show \a div b = 0\ by simp qed with \b \ 0\ show ?thesis by simp qed lemma div_mult1_eq: \(a * b) div c = a * (b div c) + a * (b mod c) div c\ proof - have *: \(a * b) mod c + (a * (c * (b div c)) + c * (a * (b mod c) div c)) = a * b\ (is \?A + (?B + ?C) = _\) proof - have \?A = a * (b mod c) mod c\ by (simp add: mod_mult_right_eq) then have \?C + ?A = a * (b mod c)\ by (simp add: mult_div_mod_eq) then have \?B + (?C + ?A) = a * (c * (b div c) + (b mod c))\ by (simp add: algebra_simps) also have \\ = a * b\ by (simp add: mult_div_mod_eq) finally show ?thesis by (simp add: algebra_simps) qed have \((a * b) div c, (a * b) mod c) = (a * (b div c) + a * (b mod c) div c, (a * b) mod c)\ proof (cases c \a * (b div c) + a * (b mod c) div c\ \(a * b) mod c\ \a * b\ rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed lemma div_add1_eq: \(a + b) div c = a div c + b div c + (a mod c + b mod c) div c\ proof - have *: \(a + b) mod c + (c * (a div c) + (c * (b div c) + c * ((a mod c + b mod c) div c))) = a + b\ (is \?A + (?B + (?C + ?D)) = _\) proof - have \?A + (?B + (?C + ?D)) = ?A + ?D + (?B + ?C)\ by (simp add: ac_simps) also have \?A + ?D = (a mod c + b mod c) mod c + ?D\ by (simp add: mod_add_eq) also have \\ = a mod c + b mod c\ by (simp add: mod_mult_div_eq) finally have \?A + (?B + (?C + ?D)) = (a mod c + ?B) + (b mod c + ?C)\ by (simp add: ac_simps) then show ?thesis by (simp add: mod_mult_div_eq) qed have \((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)\ proof (cases c \a div c + b div c + (a mod c + b mod c) div c\ \(a + b) mod c\ \a + b\ rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: \nat \ nat\ where [simp]: \normalize = (id :: nat \ nat)\ definition unit_factor_nat :: \nat \ nat\ where \unit_factor n = of_bool (n > 0)\ for n :: nat lemma unit_factor_simps [simp]: \unit_factor 0 = (0::nat)\ \unit_factor (Suc n) = 1\ by (simp_all add: unit_factor_nat_def) definition divide_nat :: \nat \ nat \ nat\ where \m div n = (if n = 0 then 0 else Max {k. k * n \ m})\ for m n :: nat instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: \nat \ nat\ where [simp]: \euclidean_size_nat = id\ definition division_segment_nat :: \nat \ nat\ where [simp]: \division_segment n = 1\ for n :: nat definition modulo_nat :: \nat \ nat \ nat\ where \m mod n = m - (m div n * n)\ for m n :: nat instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end lemma euclidean_relation_natI [case_names by0 divides euclidean_relation]: \(m div n, m mod n) = (q, r)\ if by0: \n = 0 \ q = 0 \ r = m\ and divides: \n > 0 \ n dvd m \ r = 0 \ m = q * n\ and euclidean_relation: \n > 0 \ \ n dvd m \ r < n \ m = q * n + r\ for m n q r :: nat by (rule euclidean_relationI) (use that in simp_all) lemma div_nat_eqI: \m div n = q\ if \n * q \ m\ and \m < n * Suc q\ for m n q :: nat proof - have \(m div n, m mod n) = (q, m - n * q)\ proof (cases n q \m - n * q\ m rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from \n dvd m\ obtain s where \m = n * s\ .. with \n > 0\ that have \s < Suc q\ by (simp only: mult_less_cancel1) with \m = n * s\ \n > 0\ that have \q = s\ by simp with \m = n * s\ show ?case by (simp add: ac_simps) next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed lemma mod_nat_eqI: \m mod n = r\ if \r < n\ and \r \ m\ and \n dvd m - r\ for m n r :: nat proof - have \(m div n, m mod n) = ((m - r) div n, r)\ proof (cases n \(m - r) div n\ r m rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from that dvd_minus_add [of r \m\ 1 n] have \n dvd m + (n - r)\ by simp with divides have \n dvd n - r\ by (simp add: dvd_add_right_iff) then have \n \ n - r\ by (rule dvd_imp_le) (use \r < n\ in simp) with \n > 0\ have \r = 0\ by simp with \n > 0\ that show ?case by simp next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_nat_eqI mod_nat_eqI) lemma split_div: \P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))\ (is ?div) and split_mod: \Q (m mod n) \ (n = 0 \ Q m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ Q j))\ (is ?mod) for m n :: nat proof - have *: \R (m div n) (m mod n) \ (n = 0 \ R 0 m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ R i j))\ for R by (cases \n = 0\) auto from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed declare split_div [of _ _ \numeral n\, linarith_split] for n declare split_mod [of _ _ \numeral n\, linarith_split] for n lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all lemma div_mult2_eq: \m div (n * q) = (m div n) div q\ (is ?Q) and mod_mult2_eq: \m mod (n * q) = n * (m div n mod q) + m mod n\ (is ?R) for m n q :: nat proof - have \(m div (n * q), m mod (n * q)) = ((m div n) div q, n * (m div n mod q) + m mod n)\ proof (cases \n * q\ \(m div n) div q\ \n * (m div n mod q) + m mod n\ m rule: euclidean_relation_natI) case by0 then show ?case by auto next case divides from \n * q dvd m\ obtain t where \m = n * q * t\ .. with \n * q > 0\ show ?case by (simp add: algebra_simps) next case euclidean_relation then have \n > 0\ \q > 0\ by simp_all from \n > 0\ have \m mod n < n\ by (rule mod_less_divisor) from \q > 0\ have \m div n mod q < q\ by (rule mod_less_divisor) then obtain s where \q = Suc (m div n mod q + s)\ by (blast dest: less_imp_Suc_add) moreover have \m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)\ using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately have \m mod n + n * (m div n mod q) < n * q\ by simp then show ?case by (simp add: algebra_simps flip: add_mult_distrib2) qed then show ?Q and ?R by simp_all qed lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed lemma div_less_iff_less_mult: \m div q < n \ m < n * q\ (is \?P \ ?Q\) if \q > 0\ for m n q :: nat proof assume ?Q then show ?P by (rule less_mult_imp_div_less) next assume ?P then obtain h where \n = Suc (m div q + h)\ using less_natE by blast moreover have \m < m + (Suc h * q - m mod q)\ using that by (simp add: trans_less_add1) ultimately show ?Q by (simp add: algebra_simps flip: minus_mod_eq_mult_div) qed lemma less_eq_div_iff_mult_less_eq: \m \ n div q \ m * q \ n\ if \q > 0\ for m n q :: nat using div_less_iff_less_mult [of q n m] that by auto -text \A fact for the mutilated chess board\ +lemma div_Suc: + \Suc m div n = (if Suc m mod n = 0 then Suc (m div n) else m div n)\ (is "_ = ?rhs") +proof (cases \n = 0 \ n = 1\) + case True + then show ?thesis by auto +next + case False + then have \n > 1\ + by simp + then have *: \Suc 0 div n = 0\ + by (simp add: div_eq_0_iff) + have \(m + 1) div n = ?rhs\ + proof (cases \n dvd Suc m\) + case True + then obtain q where \Suc m = n * q\ .. + then have m: \m = n * q - 1\ + by simp + have \q > 0\ by (rule ccontr) + (use \Suc m = n * q\ in simp) + from m have \m mod n = (n * q - 1) mod n\ + by simp + also have \\ = (n * q - 1 + n) mod n\ + by simp + also have \n * q - 1 + n = n * q + (n - 1)\ + using \n > 1\ \q > 0\ by (simp add: algebra_simps) + finally have \m mod n = (n - 1) mod n\ + by simp + with \n > 1\ have \m mod n = n - 1\ + by simp + with True \n > 1\ show ?thesis + by (subst div_add1_eq) auto + next + case False + have \Suc (m mod n) \ n\ + proof (rule ccontr) + assume \\ Suc (m mod n) \ n\ + then have \m mod n = n - 1\ + by simp + with \n > 1\ have \(m + 1) mod n = 0\ + by (subst mod_add_left_eq [symmetric]) simp + then have \n dvd Suc m\ + by auto + with False show False .. + qed + moreover have \Suc (m mod n) \ n\ + using \n > 1\ by (simp add: Suc_le_eq) + ultimately have \Suc (m mod n) < n\ + by simp + with False \n > 1\ show ?thesis + by (subst div_add1_eq) (auto simp add: div_eq_0_iff mod_greater_zero_iff_not_dvd) + qed + then show ?thesis + by simp +qed lemma mod_Suc: - "Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))" (is "_ = ?rhs") + \Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))\ (is "_ = ?rhs") proof (cases "n = 0") case True then show ?thesis by simp next case False have "Suc m mod n = Suc (m mod n) mod n" by (simp add: mod_simps) also have "\ = ?rhs" using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq) finally show ?thesis . qed lemma Suc_times_mod_eq: "Suc (m * n) mod m = 1" if "Suc 0 < m" using that by (simp add: mod_Suc) lemma Suc_times_numeral_mod_eq [simp]: "Suc (numeral k * n) mod numeral k = 1" if "numeral k \ (1::nat)" by (rule Suc_times_mod_eq) (use that in simp) lemma Suc_div_le_mono [simp]: "m div n \ Suc m div n" by (simp add: div_le_mono) text \These lemmas collapse some needless occurrences of Suc: at least three Sucs, since two and fewer are rewritten back to Suc again! We already have some rules to simplify operands smaller than 3.\ lemma div_Suc_eq_div_add3 [simp]: "m div Suc (Suc (Suc n)) = m div (3 + n)" by (simp add: Suc3_eq_add_3) lemma mod_Suc_eq_mod_add3 [simp]: "m mod Suc (Suc (Suc n)) = m mod (3 + n)" by (simp add: Suc3_eq_add_3) lemma Suc_div_eq_add3_div: "Suc (Suc (Suc m)) div n = (3 + m) div n" by (simp add: Suc3_eq_add_3) lemma Suc_mod_eq_add3_mod: "Suc (Suc (Suc m)) mod n = (3 + m) mod n" by (simp add: Suc3_eq_add_3) lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v lemma (in field_char_0) of_nat_div: "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)" proof - have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)" unfolding of_nat_add by (cases "n = 0") simp_all then show ?thesis by simp qed text \An ``induction'' law for modulus arithmetic.\ lemma mod_induct [consumes 3, case_names step]: "P m" if "P n" and "n < p" and "m < p" and step: "\n. n < p \ P n \ P (Suc n mod p)" using \m < p\ proof (induct m) case 0 show ?case proof (rule ccontr) assume "\ P 0" from \n < p\ have "0 < p" by simp from \n < p\ obtain m where "0 < m" and "p = n + m" by (blast dest: less_imp_add_positive) with \P n\ have "P (p - m)" by simp moreover have "\ P (p - m)" using \0 < m\ proof (induct m) case 0 then show ?case by simp next case (Suc m) show ?case proof assume P: "P (p - Suc m)" with \\ P 0\ have "Suc m < p" by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma funpow_mod_eq: \<^marker>\contributor \Lars Noschinski\\ \(f ^^ (m mod n)) x = (f ^^ m) x\ if \(f ^^ n) x = x\ proof - have \(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\ by simp also have \\ = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\ by (simp only: funpow_add funpow_mult ac_simps) simp also have \((f ^^ n) ^^ q) x = x\ for q by (induction q) (use \(f ^^ n) x = x\ in simp_all) finally show ?thesis by simp qed lemma mod_eq_dvd_iff_nat: \m mod q = n mod q \ q dvd m - n\ (is \?P \ ?Q\) if \m \ n\ for m n q :: nat proof assume ?Q then obtain s where \m - n = q * s\ .. with that have \m = q * s + n\ by simp then show ?P by simp next assume ?P have \m - n = m div q * q + m mod q - (n div q * q + n mod q)\ by simp also have \\ = q * (m div q - n div q)\ by (simp only: algebra_simps \?P\) finally show ?Q .. qed +lemma mod_eq_iff_dvd_symdiff_nat: + \m mod q = n mod q \ q dvd nat \int m - int n\\ + by (auto simp add: abs_if mod_eq_dvd_iff_nat nat_diff_distrib dest: sym intro: sym) + lemma mod_eq_nat1E: fixes m n q :: nat assumes "m mod q = n mod q" and "m \ n" obtains s where "m = n + q * s" proof - from assms have "q dvd m - n" by (simp add: mod_eq_dvd_iff_nat) then obtain s where "m - n = q * s" .. with \m \ n\ have "m = n + q * s" by simp with that show thesis . qed lemma mod_eq_nat2E: fixes m n q :: nat assumes "m mod q = n mod q" and "n \ m" obtains s where "n = m + q * s" using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" (is "?lhs = ?rhs") proof assume H: "x mod n = y mod n" { assume xy: "x \ y" from H have th: "y mod n = x mod n" by simp from mod_eq_nat1E [OF th xy] obtain q where "y = x + n * q" . then have "x + n * q = y + n * 0" by simp then have "\q1 q2. x + n * q1 = y + n * q2" by blast } moreover { assume xy: "y \ x" from mod_eq_nat1E [OF H xy] obtain q where "x = y + n * q" . then have "x + n * 0 = y + n * q" by simp then have "\q1 q2. x + n * q1 = y + n * q2" by blast } ultimately show ?rhs using linear[of x y] by blast next assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp thus ?lhs by simp qed subsection \Elementary euclidean division on \<^typ>\int\\ subsubsection \Basic instantiation\ instantiation int :: "{normalization_semidom, idom_modulo}" begin definition normalize_int :: \int \ int\ where [simp]: \normalize = (abs :: int \ int)\ definition unit_factor_int :: \int \ int\ where [simp]: \unit_factor = (sgn :: int \ int)\ definition divide_int :: \int \ int \ int\ where \k div l = (sgn k * sgn l * int (nat \k\ div nat \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k))\ lemma divide_int_unfold: \(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n) - of_bool ((k = 0 \ m = 0) \ l \ 0 \ n \ 0 \ sgn k \ sgn l \ \ n dvd m))\ by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps) definition modulo_int :: \int \ int \ int\ where \k mod l = sgn k * int (nat \k\ mod nat \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ lemma modulo_int_unfold: \(sgn k * int m) mod (sgn l * int n) = sgn k * int (m mod (of_bool (l \ 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \ m = 0) \ sgn k \ sgn l \ \ n dvd m)\ by (auto simp add: modulo_int_def sgn_mult abs_mult) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) next fix k l :: int obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end subsubsection \Algebraic foundations\ lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all subsubsection \Basic conversions\ lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (auto simp add: divide_int_def) lemma div_eq_div_abs: \k div l = sgn k * sgn l * (\k\ div \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: divide_int_def [of k l] div_abs_eq_div_nat) lemma div_abs_eq: \\k\ div \l\ = sgn k * sgn l * (k div l + of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (simp add: div_eq_div_abs [of k l] ac_simps) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma mod_eq_mod_abs: \k mod l = sgn k * (\k\ mod \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat) lemma mod_abs_eq: \\k\ mod \l\ = sgn k * (k mod l - l * of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (auto simp: mod_eq_mod_abs [of k l]) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" using assms by (simp add: sgn_mult abs_mult sgn_0_0 divide_int_def [of "sgn v * \k\" "sgn v * \l\"] flip: div_abs_eq_div_nat) lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" using assms by (auto simp add: div_abs_eq) lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" using assms by (auto simp add: div_abs_eq ac_simps) lemma div_noneq_sgn_abs: fixes k l :: int assumes "l \ 0" assumes "sgn k \ sgn l" shows "k div l = - (\k\ div \l\) - of_bool (\ l dvd k)" using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp) subsubsection \Euclidean division\ instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) (simp add: sgn_0_0) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: * q l divide_int_unfold) (auto simp add: sgn_mult ac_simps) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \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\) end lemma euclidean_relation_intI [case_names by0 divides euclidean_relation]: \(k div l, k mod l) = (q, r)\ if by0': \l = 0 \ q = 0 \ r = k\ and divides': \l \ 0 \ l dvd k \ r = 0 \ k = q * l\ and euclidean_relation': \l \ 0 \ \ l dvd k \ sgn r = sgn l \ \r\ < \l\ \ k = q * l + r\ for k l :: int proof (cases l q r k rule: euclidean_relationI) case by0 then show ?case by (rule by0') next case divides then show ?case by (rule divides') next case euclidean_relation with euclidean_relation' have \sgn r = sgn l\ \\r\ < \l\\ \k = q * l + r\ by simp_all from \sgn r = sgn l\ \l \ 0\ have \division_segment r = division_segment l\ by (simp add: division_segment_int_def sgn_if split: if_splits) with \\r\ < \l\\ \k = q * l + r\ show ?case by simp qed subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ proof (cases \m = 0 \ n = 0\) case True then show ?thesis by auto next case False then have \m > 0\ \n > 0\ by simp_all show ?thesis proof (cases \of_nat m * of_nat n dvd a\) case True then obtain b where \a = (of_nat m * of_nat n) * b\ .. then have \a = of_nat m * (of_nat n * b)\ by (simp add: ac_simps) then show ?thesis by simp next case False define q where \q = a div (of_nat m * of_nat n)\ define r where \r = a mod (of_nat m * of_nat n)\ from \m > 0\ \n > 0\ \\ of_nat m * of_nat n dvd a\ r_def have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod) with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with \m > 0\ \n > 0\ r_def have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with \m > 0\ \n > 0\ q_def have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp moreover have \a = q * (of_nat m * of_nat n) + r\ by (simp add: q_def r_def div_mult_mod_eq) ultimately show \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ using q_def [symmetric] div_plus_div_distrib_dvd_right [of \of_nat m\ \q * (of_nat m * of_nat n)\ r] by (simp add: ac_simps) qed qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np) subsection \More on euclidean division on \<^typ>\int\\ subsubsection \Trivial reduction steps\ lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_pos_neg_trivial: \k div l = - 1\ (is ?Q) and mod_pos_neg_trivial: \k mod l = k + l\ (is ?R) if \0 < k\ and \k + l \ 0\ for k l :: int proof - from that have \l < 0\ by simp have \(k div l, k mod l) = (- 1, k + l)\ proof (cases l \- 1 :: int\ \k + l\ k rule: euclidean_relation_intI) case by0 with \l < 0\ show ?case by simp next case divides from \l dvd k\ obtain j where \k = l * j\ .. with \l < 0\ \0 < k\ have \j < 0\ by (simp add: zero_less_mult_iff) moreover from \k + l \ 0\ \k = l * j\ have \l * (j + 1) \ 0\ by (simp add: algebra_simps) with \l < 0\ have \j + 1 \ 0\ by (simp add: mult_le_0_iff) with \j < 0\ have \j = - 1\ by simp with \k = l * j\ show ?case by simp next case euclidean_relation with \k + l \ 0\ have \k + l < 0\ by (auto simp add: less_le add_eq_0_iff) with \0 < k\ show ?case by simp qed then show ?Q and ?R by simp_all qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ subsubsection \More uniqueness rules\ lemma fixes a b q r :: int assumes \a = b * q + r\ \0 \ r\ \r < b\ shows int_div_pos_eq: \a div b = q\ (is ?Q) and int_mod_pos_eq: \a mod b = r\ (is ?R) proof - from assms have \(a div b, a mod b) = (q, r)\ by (cases b q r a rule: euclidean_relation_intI) (auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le) then show ?Q and ?R by simp_all qed lemma int_div_neg_eq: \a div b = q\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int using that int_div_pos_eq [of a \- b\ \- q\ \- r\] by simp_all lemma int_mod_neg_eq: \a mod b = r\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int using that int_div_neg_eq [of a b q r] by simp subsubsection \Laws for unary minus\ lemma zmod_zminus1_not_zero: fixes k l :: int shows "- k mod l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zmod_zminus2_not_zero: fixes k l :: int shows "k mod - l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zdiv_zminus1_eq_if: \(- a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that sgn_not_eq_imp [of b \- a\] by (cases \a = 0\) (auto simp add: div_eq_div_abs [of \- a\ b] div_eq_div_abs [of a b] sgn_eq_0_iff) lemma zdiv_zminus2_eq_if: \a div (- b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that by (auto simp add: zdiv_zminus1_eq_if div_minus_right) lemma zmod_zminus1_eq_if: \(- a) mod b = (if a mod b = 0 then 0 else b - (a mod b))\ for a b :: int by (cases \b = 0\) (auto simp flip: minus_div_mult_eq_mod simp add: zdiv_zminus1_eq_if algebra_simps) lemma zmod_zminus2_eq_if: \a mod (- b) = (if a mod b = 0 then 0 else (a mod b) - b)\ for a b :: int by (auto simp add: zmod_zminus1_eq_if mod_minus_right) subsubsection \Borders\ lemma pos_mod_bound [simp]: "k mod l < l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_pos) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_neg) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: sgn_1_pos) qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp moreover have \int (m mod n) \ int n\ using \Suc q = n\ by simp then have \sgn s * int (m mod n) \ int n\ by (cases s \0::int\ rule: linorder_cases) simp_all ultimately show ?thesis by (simp only: modulo_int_unfold) auto qed subsubsection \Splitting Rules for div and mod\ lemma split_zdiv: \P (n div k) \ (k = 0 \ P 0) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P i)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P i))\ (is ?div) and split_zmod: \Q (n mod k) \ (k = 0 \ Q n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ Q j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ Q j))\ (is ?mod) for n k :: int proof - have *: \R (n div k) (n mod k) \ (k = 0 \ R 0 n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ R i j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ R i j))\ for R by (cases \k = 0\) (auto simp add: linorder_class.neq_iff) from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed text \Enable (lin)arith to deal with \<^const>\divide\ and \<^const>\modulo\ when these are applied to some constant that is of the form \<^term>\numeral k\:\ declare split_zdiv [of _ _ \numeral n\, linarith_split] for n declare split_zdiv [of _ _ \- numeral n\, linarith_split] for n declare split_zmod [of _ _ \numeral n\, linarith_split] for n declare split_zmod [of _ _ \- numeral n\, linarith_split] for n lemma zdiv_eq_0_iff: "i div k = 0 \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" (is "?L = ?R") for i k :: int proof assume ?L moreover have "?L \ ?R" by (rule split_zdiv [THEN iffD2]) simp ultimately show ?R by blast next assume ?R then show ?L by auto qed lemma zmod_trivial_iff: fixes i k :: int shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" proof - have "i mod k = i \ i div k = 0" using div_mult_mod_eq [of i k] by safe auto with zdiv_eq_0_iff show ?thesis by simp qed subsubsection \Algebraic rewrites\ lemma zdiv_zmult2_eq: \a div (b * c) = (a div b) div c\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using div_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using div_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma zdiv_zmult2_eq': \k div (l * j) = ((sgn j * k) div l) div \j\\ for k l j :: int proof - have \k div (l * j) = (sgn j * k) div (sgn j * (l * j))\ by (simp add: sgn_0_0) also have \sgn j * (l * j) = l * \j\\ by (simp add: mult.left_commute [of _ l] abs_sgn) (simp add: ac_simps) also have \(sgn j * k) div (l * \j\) = ((sgn j * k) div l) div \j\\ by (simp add: zdiv_zmult2_eq) finally show ?thesis . qed lemma zmod_zmult2_eq: \a mod (b * c) = b * (a div b mod c) + a mod b\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using mod_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using mod_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma half_nonnegative_int_iff [simp]: \k div 2 \ 0 \ k \ 0\ for k :: int by auto lemma half_negative_int_iff [simp]: \k div 2 < 0 \ k < 0\ for k :: int by auto subsubsection \Distributive laws for conversions.\ lemma zdiv_int: "int (a div b) = int a div int b" by (fact of_nat_div) lemma zmod_int: "int (a mod b) = int a mod int b" by (fact of_nat_mod) lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) subsubsection \Monotonicity in the First Argument (Dividend)\ lemma zdiv_mono1: \a div b \ a' div b\ if \a \ a'\ \0 < b\ for a b b' :: int proof - from \a \ a'\ have \b * (a div b) + a mod b \ b * (a' div b) + a' mod b\ by simp then have \b * (a div b) \ (a' mod b - a mod b) + b * (a' div b)\ by (simp add: algebra_simps) moreover have \a' mod b < b + a mod b\ by (rule less_le_trans [of _ b]) (use \0 < b\ in simp_all) ultimately have \b * (a div b) < b * (1 + a' div b)\ by (simp add: distrib_left) with \0 < b\ have \a div b < 1 + a' div b\ by (simp add: mult_less_cancel_left) then show ?thesis by simp qed lemma zdiv_mono1_neg: \a' div b \ a div b\ if \a \ a'\ \b < 0\ for a a' b :: int using that zdiv_mono1 [of \- a'\ \- a\ \- b\] by simp subsubsection \Monotonicity in the Second Argument (Divisor)\ lemma zdiv_mono2: \a div b \ a div b'\ if \0 \ a\ \0 < b'\ \b' \ b\ for a b b' :: int proof - define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ then have *: \b * q + r = b' * q' + r'\ \0 \ b' * q' + r'\ \r' < b'\ \0 \ r\ \0 < b'\ \b' \ b\ using that by simp_all have \0 < b' * (q' + 1)\ using * by (simp add: distrib_left) with * have \0 \ q'\ by (simp add: zero_less_mult_iff) moreover have \b * q = r' - r + b' * q'\ using * by linarith ultimately have \b * q < b * (q' + 1)\ using mult_right_mono * unfolding distrib_left by fastforce with * have \q \ q'\ by (simp add: mult_less_cancel_left_pos) with ** show ?thesis by simp qed lemma zdiv_mono2_neg: \a div b' \ a div b\ if \a < 0\ \0 < b'\ \b' \ b\ for a b b' :: int proof - define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ then have *: \b * q + r = b' * q' + r'\ \b' * q' + r' < 0\ \r < b\ \0 \ r'\ \0 < b'\ \b' \ b\ using that by simp_all have \b' * q' < 0\ using * by linarith with * have \q' \ 0\ by (simp add: mult_less_0_iff) have \b * q' \ b' * q'\ by (simp add: \q' \ 0\ * mult_right_mono_neg) then have "b * q' < b * (q + 1)" using * by (simp add: distrib_left) then have \q' \ q\ using * by (simp add: mult_less_cancel_left) then show ?thesis by (simp add: **) qed subsubsection \Quotients of Signs\ lemma div_eq_minus1: \0 < b \ - 1 div b = - 1\ for b :: int by (simp add: divide_int_def) lemma zmod_minus1: \0 < b \ - 1 mod b = b - 1\ for b :: int by (auto simp add: modulo_int_def) lemma minus_mod_int_eq: \- k mod l = l - 1 - (k - 1) mod l\ if \l \ 0\ for k l :: int proof (cases \l = 0\) case True then show ?thesis by simp next case False with that have \l > 0\ by simp then show ?thesis proof (cases \l dvd k\) case True then obtain j where \k = l * j\ .. moreover have \(l * j mod l - 1) mod l = l - 1\ using \l > 0\ by (simp add: zmod_minus1) then have \(l * j - 1) mod l = l - 1\ by (simp only: mod_simps) ultimately show ?thesis by simp next case False moreover have 1: \0 < k mod l\ using \0 < l\ False le_less by fastforce moreover have 2: \k mod l < 1 + l\ using \0 < l\ pos_mod_bound[of l k] by linarith from 1 2 \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ by (simp add: zmod_trivial_iff) ultimately show ?thesis by (simp only: zmod_zminus1_eq_if) (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) qed qed lemma div_neg_pos_less0: \a div b < 0\ if \a < 0\ \0 < b\ for a b :: int proof - have "a div b \ - 1 div b" using zdiv_mono1 that by auto also have "... \ -1" by (simp add: that(2) div_eq_minus1) finally show ?thesis by force qed lemma div_nonneg_neg_le0: \a div b \ 0\ if \0 \ a\ \b < 0\ for a b :: int using that by (auto dest: zdiv_mono1_neg) lemma div_nonpos_pos_le0: \a div b \ 0\ if \a \ 0\ \0 < b\ for a b :: int using that by (auto dest: zdiv_mono1) text\Now for some equivalences of the form \a div b >=< 0 \ \\ conditional upon the sign of \a\ or \b\. There are many more. They should all be simp rules unless that causes too much search.\ lemma pos_imp_zdiv_nonneg_iff: \0 \ a div b \ 0 \ a\ if \0 < b\ for a b :: int proof assume \0 \ a div b\ show \0 \ a\ proof (rule ccontr) assume \\ 0 \ a\ then have \a < 0\ by simp then have \a div b < 0\ using that by (rule div_neg_pos_less0) with \0 \ a div b\ show False by simp qed next assume "0 \ a" then have "0 div b \ a div b" using zdiv_mono1 that by blast then show "0 \ a div b" by auto qed lemma neg_imp_zdiv_nonneg_iff: \0 \ a div b \ a \ 0\ if \b < 0\ for a b :: int using that pos_imp_zdiv_nonneg_iff [of \- b\ \- a\] by simp lemma pos_imp_zdiv_pos_iff: \0 < (i::int) div k \ k \ i\ if \0 < k\ for i k :: int using that pos_imp_zdiv_nonneg_iff [of k i] zdiv_eq_0_iff [of i k] by arith lemma pos_imp_zdiv_neg_iff: \a div b < 0 \ a < 0\ if \0 < b\ for a b :: int \ \But not \<^prop>\a div b \ 0 \ a \ 0\; consider \<^prop>\a = 1\, \<^prop>\b = 2\ when \<^prop>\a div b = 0\.\ using that by (simp add: pos_imp_zdiv_nonneg_iff flip: linorder_not_le) lemma neg_imp_zdiv_neg_iff: \ \But not \<^prop>\a div b \ 0 \ 0 \ a\; consider \<^prop>\a = - 1\, \<^prop>\b = - 2\ when \<^prop>\a div b = 0\.\ \a div b < 0 \ 0 < a\ if \b < 0\ for a b :: int using that by (simp add: neg_imp_zdiv_nonneg_iff flip: linorder_not_le) lemma nonneg1_imp_zdiv_pos_iff: \a div b > 0 \ a \ b \ b > 0\ if \0 \ a\ for a b :: int proof - have "0 < a div b \ b \ a" using div_pos_pos_trivial[of a b] that by arith moreover have "0 < a div b \ b > 0" using that div_nonneg_neg_le0[of a b] by (cases "b=0"; force) moreover have "b \ a \ 0 < b \ 0 < a div b" using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp ultimately show ?thesis by blast qed lemma zmod_le_nonneg_dividend: \m mod k \ m\ if \(m::int) \ 0\ for m k :: int proof - from that have \m > 0 \ m = 0\ by auto then show ?thesis proof assume \m = 0\ then show ?thesis by simp next assume \m > 0\ then show ?thesis proof (cases k \0::int\ rule: linorder_cases) case less moreover define l where \l = - k\ ultimately have \l > 0\ by simp with \m > 0\ have \int (nat m mod nat l) \ m\ by (simp flip: le_nat_iff) then have \int (nat m mod nat l) - l \ m\ using \l > 0\ by simp with \m > 0\ \l > 0\ show ?thesis by (simp add: modulo_int_def l_def flip: le_nat_iff) qed (simp_all add: modulo_int_def flip: le_nat_iff) qed qed lemma sgn_div_eq_sgn_mult: \sgn (k div l) = of_bool (k div l \ 0) * sgn (k * l)\ for k l :: int proof (cases \k div l = 0\) case True then show ?thesis by simp next case False have \0 \ \k\ div \l\\ by (cases \l = 0\) (simp_all add: pos_imp_zdiv_nonneg_iff) then have \\k\ div \l\ \ 0 \ 0 < \k\ div \l\\ by (simp add: less_le) also have \\ \ \k\ \ \l\\ using False nonneg1_imp_zdiv_pos_iff by auto finally have *: \\k\ div \l\ \ 0 \ \l\ \ \k\\ . show ?thesis using \0 \ \k\ div \l\\ False by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l] sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp) qed subsubsection \Further properties\ lemma div_int_pos_iff: "k div l \ 0 \ k = 0 \ l = 0 \ k \ 0 \ l \ 0 \ k < 0 \ l < 0" for k l :: int proof (cases "k = 0 \ l = 0") case False then have *: "k \ 0" "l \ 0" by auto then have "0 \ k div l \ \ k < 0 \ 0 \ l" by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) then show ?thesis using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) qed auto lemma mod_int_pos_iff: \k mod l \ 0 \ l dvd k \ l = 0 \ k \ 0 \ l > 0\ for k l :: int proof (cases "l > 0") case False then show ?thesis by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \auto simp add: le_less not_less\) qed auto lemma abs_div: \\x div y\ = \x\ div \y\\ if \y dvd x\ for x y :: int using that by (cases \y = 0\) (auto simp add: abs_mult) lemma int_power_div_base: \<^marker>\contributor \Matthias Daum\\ \k ^ m div k = k ^ (m - Suc 0)\ if \0 < m\ \0 < k\ for k :: int using that by (cases m) simp_all lemma int_div_less_self: \<^marker>\contributor \Matthias Daum\\ \x div k < x\ if \0 < x\ \1 < k\ for x k :: int proof - from that have \nat (x div k) = nat x div nat k\ by (simp add: nat_div_distrib) also from that have \nat x div nat k < nat x\ by simp finally show ?thesis by simp qed subsubsection \Computing \div\ and \mod\ by shifting\ lemma div_pos_geq: \k div l = (k - l) div l + 1\ if \0 < l\ \l \ k\ for k l :: int proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with that show ?thesis by (simp add: div_add_self2) qed lemma mod_pos_geq: \k mod l = (k - l) mod l\ if \0 < l\ \l \ k\ for k l :: int proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with that show ?thesis by simp qed lemma pos_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = b div a\ (is ?Q) and pos_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\ (is ?R) if \0 \ a\ for a b :: int proof - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\ proof (cases \2 * a\ \b div a\ \1 + 2 * (b mod a)\ \1 + 2 * b\ rule: euclidean_relation_intI) case by0 then show ?case by simp next case divides have \2 dvd (2 * a)\ by simp then have \2 dvd (1 + 2 * b)\ using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) then have \2 dvd (1 + b * 2)\ by (simp add: ac_simps) then have \is_unit (2 :: int)\ by simp then show ?case by simp next case euclidean_relation with that have \a > 0\ by simp moreover have \b mod a < a\ using \a > 0\ by simp then have \1 + 2 * (b mod a) < 2 * a\ by simp moreover have \2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\ by (simp only: algebra_simps) moreover have \0 \ 2 * (b mod a)\ using \a > 0\ by simp ultimately show ?case by (simp add: algebra_simps) qed then show ?Q and ?R by simp_all qed lemma neg_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = (b + 1) div a\ (is ?Q) and neg_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\ (is ?R) if \a \ 0\ for a b :: int proof - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\ proof (cases \2 * a\ \(b + 1) div a\ \2 * ((b + 1) mod a) - 1\ \1 + 2 * b\ rule: euclidean_relation_intI) case by0 then show ?case by simp next case divides have \2 dvd (2 * a)\ by simp then have \2 dvd (1 + 2 * b)\ using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) then have \2 dvd (1 + b * 2)\ by (simp add: ac_simps) then have \is_unit (2 :: int)\ by simp then show ?case by simp next case euclidean_relation with that have \a < 0\ by simp moreover have \(b + 1) mod a > a\ using \a < 0\ by simp then have \2 * ((b + 1) mod a) > 1 + 2 * a\ by simp moreover have \((1 + b) mod a) \ 0\ using \a < 0\ by simp then have \2 * ((1 + b) mod a) \ 0\ by simp moreover have \2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) = 2 * ((1 + b) div a * a + (1 + b) mod a)\ by (simp only: algebra_simps) ultimately show ?case by (simp add: algebra_simps sgn_mult abs_mult) qed then show ?Q and ?R by simp_all qed lemma zdiv_numeral_Bit0 [simp]: \numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = numeral v div (numeral w :: int)\ unfolding numeral.simps unfolding mult_2 [symmetric] by (rule div_mult_mult1) simp lemma zdiv_numeral_Bit1 [simp]: \numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = (numeral v div (numeral w :: int))\ unfolding numeral.simps unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zdiv_mult_2) simp lemma zmod_numeral_Bit0 [simp]: \numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = (2::int) * (numeral v mod numeral w)\ unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] by (rule mod_mult_mult1) lemma zmod_numeral_Bit1 [simp]: \numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = 2 * (numeral v mod numeral w) + (1::int)\ unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zmod_mult_2) simp subsection \Generic symbolic computations\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. \ class unique_euclidean_semiring_with_nat_division = unique_euclidean_semiring_with_nat + fixes divmod :: \num \ num \ 'a \ 'a\ and divmod_step :: \'a \ 'a \ 'a \ 'a \ 'a\ \ \ These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ assumes divmod_def: \divmod m n = (numeral m div numeral n, numeral m mod numeral n)\ and divmod_step_def [simp]: \divmod_step l (q, r) = (if euclidean_size l \ euclidean_size r then (2 * q + 1, r - l) else (2 * q, r))\ \ \ This is a formulation of one step (referring to one digit position) in school-method division: compare the dividend at the current digit position with the remainder from previous division steps and evaluate accordingly.\ begin lemma fst_divmod: \fst (divmod m n) = numeral m div numeral n\ by (simp add: divmod_def) lemma snd_divmod: \snd (divmod m n) = numeral m mod numeral n\ by (simp add: divmod_def) text \ Following a formulation of school-method division. If the divisor is smaller than the dividend, terminate. If not, shift the dividend to the right until termination occurs and then reiterate single division steps in the opposite direction. \ lemma divmod_divmod_step: \divmod m n = (if m < n then (0, numeral m) else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\ proof (cases \m < n\) case True then show ?thesis by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod) next case False define r s t where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ \t = 2 * s\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 n) = of_nat t\ and \\ s \ r mod s\ by (simp_all add: not_le) have t: \2 * (r div t) = r div s - r div s mod 2\ \r mod t = s * (r div s mod 2) + r mod s\ by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Division.div_mult2_eq \t = 2 * s\) (use mod_mult2_eq [of r s 2] in \simp add: ac_simps \t = 2 * s\\) have rs: \r div s mod 2 = 0 \ r div s mod 2 = Suc 0\ by auto from \\ s \ r mod s\ have \s \ r mod t \ r div s = Suc (2 * (r div t)) \ r mod s = r mod t - s\ using rs by (auto simp add: t) moreover have \r mod t < s \ r div s = 2 * (r div t) \ r mod s = r mod t\ using rs by (auto simp add: t) ultimately show ?thesis by (simp add: divmod_def prod_eq_iff split_def Let_def not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *) (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff) qed text \The division rewrite proper -- first, trivial results involving \1\\ lemma divmod_trivial [simp]: "divmod m Num.One = (numeral m, 0)" "divmod num.One (num.Bit0 n) = (0, Numeral1)" "divmod num.One (num.Bit1 n) = (0, Numeral1)" using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) text \Division by an even number is a right-shift\ lemma divmod_cancel [simp]: \divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r))\ (is ?P) \divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r + 1))\ (is ?Q) proof - define r s where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 m) = of_nat (2 * r)\ \numeral (num.Bit0 n) = of_nat (2 * s)\ \numeral (num.Bit1 m) = of_nat (Suc (2 * r))\ by simp_all have **: \Suc (2 * r) div 2 = r\ by simp show ?P and ?Q by (simp_all add: divmod_def *) (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc add: Euclidean_Division.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **) qed text \The really hard work\ lemma divmod_steps [simp]: "divmod (num.Bit0 m) (num.Bit1 n) = (if m \ n then (0, numeral (num.Bit0 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit0 m) (num.Bit0 (num.Bit1 n))))" "divmod (num.Bit1 m) (num.Bit1 n) = (if m < n then (0, numeral (num.Bit1 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps text \Special case: divisibility\ definition divides_aux :: "'a \ 'a \ bool" where "divides_aux qr \ snd qr = 0" lemma divides_aux_eq [simp]: "divides_aux (q, r) \ r = 0" by (simp add: divides_aux_def) lemma dvd_numeral_simp [simp]: "numeral m dvd numeral n \ divides_aux (divmod n m)" by (simp add: divmod_def mod_eq_0_iff_dvd) text \Generic computation of quotient and remainder\ lemma numeral_div_numeral [simp]: "numeral k div numeral l = fst (divmod k l)" by (simp add: fst_divmod) lemma numeral_mod_numeral [simp]: "numeral k mod numeral l = snd (divmod k l)" by (simp add: snd_divmod) lemma one_div_numeral [simp]: "1 div numeral n = fst (divmod num.One n)" by (simp add: fst_divmod) lemma one_mod_numeral [simp]: "1 mod numeral n = snd (divmod num.One n)" by (simp add: snd_divmod) end instantiation nat :: unique_euclidean_semiring_with_nat_division begin definition divmod_nat :: "num \ num \ nat \ nat" where divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_nat :: "nat \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ l then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (simp_all add: divmod'_nat_def divmod_step_nat_def) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: \Suc 0 div numeral Num.One = 1\ \Suc 0 div numeral (Num.Bit0 n) = 0\ \Suc 0 div numeral (Num.Bit1 n) = 0\ by simp_all lemma Suc_0_mod_numeral [simp]: \Suc 0 mod numeral Num.One = 0\ \Suc 0 mod numeral (Num.Bit0 n) = 1\ \Suc 0 mod numeral (Num.Bit1 n) = 1\ by simp_all instantiation int :: unique_euclidean_semiring_with_nat_division begin definition divmod_int :: "num \ num \ int \ int" where "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_int :: "int \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if \l\ \ \r\ then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (auto simp add: divmod_int_def divmod_step_int_def) end declare divmod_algorithm_code [where ?'a = int, code] context begin qualified definition adjust_div :: "int \ int \ int" where "adjust_div qr = (let (q, r) = qr in q + of_bool (r \ 0))" qualified lemma adjust_div_eq [simp, code]: "adjust_div (q, r) = q + of_bool (r \ 0)" by (simp add: adjust_div_def) qualified definition adjust_mod :: "num \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)" lemma minus_numeral_div_numeral [simp]: "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma minus_numeral_mod_numeral [simp]: "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma numeral_div_minus_numeral [simp]: "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma numeral_mod_minus_numeral [simp]: "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma minus_one_div_numeral [simp]: "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" using minus_numeral_div_numeral [of Num.One n] by simp lemma minus_one_mod_numeral [simp]: "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)" using minus_numeral_mod_numeral [of Num.One n] by simp lemma one_div_minus_numeral [simp]: "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" using numeral_div_minus_numeral [of Num.One n] by simp lemma one_mod_minus_numeral [simp]: "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp lemma [code]: fixes k :: int shows "k div 0 = 0" "k mod 0 = k" "0 div k = 0" "0 mod k = 0" "k div Int.Pos Num.One = k" "k mod Int.Pos Num.One = 0" "k div Int.Neg Num.One = - k" "k mod Int.Neg Num.One = 0" "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)" "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" by simp_all end lemma divmod_BitM_2_eq [simp]: \divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\ by (cases m) simp_all subsubsection \Computation by simplification\ lemma euclidean_size_nat_less_eq_iff: \euclidean_size m \ euclidean_size n \ m \ n\ for m n :: nat by simp lemma euclidean_size_int_less_eq_iff: \euclidean_size k \ euclidean_size l \ \k\ \ \l\\ for k l :: int by auto simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | "numeral a div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \ let val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\If\); fun successful_rewrite ctxt ct = let val thm = Simplifier.rewrite ctxt ct in if Thm.is_reflexive thm then NONE else SOME thm end; in fn phi => let val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral one_div_minus_numeral one_mod_minus_numeral numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral numeral_div_minus_numeral numeral_mod_minus_numeral div_minus_minus mod_minus_minus Euclidean_Division.adjust_div_eq of_bool_eq one_neq_zero numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One case_prod_beta rel_simps Euclidean_Division.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral} @ [@{lemma "0 = 0 \ True" by simp}]); fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end end \ \ \ There is space for improvement here: the calculation itself could be carried out outside the logic, and a generic simproc (simplifier setup) for generic calculation would be helpful. \ subsubsection \Code generation\ context begin qualified definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" qualified lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = divmod_nat (m - n) n in (Suc q, r))" by (simp add: divmod_nat_def prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) qualified lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by (simp_all add: divmod_nat_def) end code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Library/Numeral_Type.thy b/src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy +++ b/src/HOL/Library/Numeral_Type.thy @@ -1,604 +1,604 @@ (* Title: HOL/Library/Numeral_Type.thy Author: Brian Huffman *) section \Numeral Syntax for Types\ theory Numeral_Type imports Cardinality begin subsection \Numeral Types\ typedef num0 = "UNIV :: nat set" .. typedef num1 = "UNIV :: unit set" .. typedef 'a bit0 = "{0 ..< 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 2 * int CARD('a)}" by simp qed typedef 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 1 + 2 * int CARD('a)}" by simp qed lemma card_num0 [simp]: "CARD (num0) = 0" unfolding type_definition.card [OF type_definition_num0] by simp lemma infinite_num0: "\ finite (UNIV :: num0 set)" using card_num0[unfolded card_eq_0_iff] by simp lemma card_num1 [simp]: "CARD(num1) = 1" unfolding type_definition.card [OF type_definition_num1] by (simp only: card_UNIV_unit) lemma card_bit0 [simp]: "CARD('a bit0) = 2 * CARD('a::finite)" unfolding type_definition.card [OF type_definition_bit0] by simp lemma card_bit1 [simp]: "CARD('a bit1) = Suc (2 * CARD('a::finite))" unfolding type_definition.card [OF type_definition_bit1] by simp subsection \@{typ num1}\ instance num1 :: finite proof show "finite (UNIV::num1 set)" unfolding type_definition.univ [OF type_definition_num1] using finite by (rule finite_imageI) qed instantiation num1 :: CARD_1 begin instance proof show "CARD(num1) = 1" by auto qed end lemma num1_eq_iff: "(x::num1) = (y::num1) \ True" by (induct x, induct y) simp instantiation num1 :: "{comm_ring,comm_monoid_mult,numeral}" begin instance by standard (simp_all add: num1_eq_iff) end lemma num1_eqI: fixes a::num1 shows "a = b" by(simp add: num1_eq_iff) lemma num1_eq1 [simp]: fixes a::num1 shows "a = 1" by (rule num1_eqI) lemma forall_1[simp]: "(\i::num1. P i) \ P 1" by (metis (full_types) num1_eq_iff) lemma ex_1[simp]: "(\x::num1. P x) \ P 1" by auto (metis (full_types) num1_eq_iff) instantiation num1 :: linorder begin definition "a < b \ Rep_num1 a < Rep_num1 b" definition "a \ b \ Rep_num1 a \ Rep_num1 b" instance by intro_classes (auto simp: less_eq_num1_def less_num1_def intro: num1_eqI) end instance num1 :: wellorder by intro_classes (auto simp: less_eq_num1_def less_num1_def) instance bit0 :: (finite) card2 proof show "finite (UNIV::'a bit0 set)" unfolding type_definition.univ [OF type_definition_bit0] by simp show "2 \ CARD('a bit0)" by simp qed instance bit1 :: (finite) card2 proof show "finite (UNIV::'a bit1 set)" unfolding type_definition.univ [OF type_definition_bit1] by simp show "2 \ CARD('a bit1)" by simp qed subsection \Locales for for modular arithmetic subtypes\ locale mod_type = fixes n :: int and Rep :: "'a::{zero,one,plus,times,uminus,minus} \ int" and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus}" assumes type: "type_definition Rep Abs {0.. n" by (rule Rep_less_n [THEN order_less_imp_le]) lemma Rep_inject_sym: "x = y \ Rep x = Rep y" by (rule type_definition.Rep_inject [OF type, symmetric]) lemma Rep_inverse: "Abs (Rep x) = x" by (rule type_definition.Rep_inverse [OF type]) lemma Abs_inverse: "m \ {0.. Rep (Abs m) = m" by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" -by (simp add: Abs_inverse pos_mod_conj [OF size0]) + using size0 by (simp add: Abs_inverse) lemma Rep_Abs_0: "Rep (Abs 0) = 0" by (simp add: Abs_inverse size0) lemma Rep_0: "Rep 0 = 0" by (simp add: zero_def Rep_Abs_0) lemma Rep_Abs_1: "Rep (Abs 1) = 1" by (simp add: Abs_inverse size1) lemma Rep_1: "Rep 1 = 1" by (simp add: one_def Rep_Abs_1) lemma Rep_mod: "Rep x mod n = Rep x" apply (rule_tac x=x in type_definition.Abs_cases [OF type]) apply (simp add: type_definition.Abs_inverse [OF type]) done lemmas Rep_simps = Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1 lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)" apply (intro_classes, unfold definitions) apply (simp_all add: Rep_simps mod_simps field_simps) done end locale mod_ring = mod_type n Rep Abs for n :: int and Rep :: "'a::{comm_ring_1} \ int" and Abs :: "int \ 'a::{comm_ring_1}" begin lemma of_nat_eq: "of_nat k = Abs (int k mod n)" apply (induct k) apply (simp add: zero_def) apply (simp add: Rep_simps add_def one_def mod_simps ac_simps) done lemma of_int_eq: "of_int z = Abs (z mod n)" apply (cases z rule: int_diff_cases) apply (simp add: Rep_simps of_nat_eq diff_def mod_simps) done lemma Rep_numeral: "Rep (numeral w) = numeral w mod n" using of_int_eq [of "numeral w"] by (simp add: Rep_inject_sym Rep_Abs_mod) lemma iszero_numeral: "iszero (numeral w::'a) \ numeral w mod n = 0" by (simp add: Rep_inject_sym Rep_numeral Rep_0 iszero_def) lemma cases: assumes 1: "\z. \(x::'a) = of_int z; 0 \ z; z < n\ \ P" shows "P" apply (cases x rule: type_definition.Abs_cases [OF type]) apply (rule_tac z="y" in 1) apply (simp_all add: of_int_eq) done lemma induct: "(\z. \0 \ z; z < n\ \ P (of_int z)) \ P (x::'a)" by (cases x rule: cases) simp lemma UNIV_eq: "(UNIV :: 'a set) = Abs ` {0..Ring class instances\ text \ Unfortunately \ring_1\ instance is not possible for \<^typ>\num1\, since 0 and 1 are not distinct. \ instantiation bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus}" begin definition Abs_bit0' :: "int \ 'a bit0" where "Abs_bit0' x = Abs_bit0 (x mod int CARD('a bit0))" definition Abs_bit1' :: "int \ 'a bit1" where "Abs_bit1' x = Abs_bit1 (x mod int CARD('a bit1))" definition "0 = Abs_bit0 0" definition "1 = Abs_bit0 1" definition "x + y = Abs_bit0' (Rep_bit0 x + Rep_bit0 y)" definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)" definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)" definition "- x = Abs_bit0' (- Rep_bit0 x)" definition "0 = Abs_bit1 0" definition "1 = Abs_bit1 1" definition "x + y = Abs_bit1' (Rep_bit1 x + Rep_bit1 y)" definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)" definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)" definition "- x = Abs_bit1' (- Rep_bit1 x)" instance .. end interpretation bit0: mod_type "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" apply (rule mod_type.intro) apply (simp add: type_definition_bit0) apply (rule one_less_int_card) apply (rule zero_bit0_def) apply (rule one_bit0_def) apply (rule plus_bit0_def [unfolded Abs_bit0'_def]) apply (rule times_bit0_def [unfolded Abs_bit0'_def]) apply (rule minus_bit0_def [unfolded Abs_bit0'_def]) apply (rule uminus_bit0_def [unfolded Abs_bit0'_def]) done interpretation bit1: mod_type "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" apply (rule mod_type.intro) apply (simp add: type_definition_bit1) apply (rule one_less_int_card) apply (rule zero_bit1_def) apply (rule one_bit1_def) apply (rule plus_bit1_def [unfolded Abs_bit1'_def]) apply (rule times_bit1_def [unfolded Abs_bit1'_def]) apply (rule minus_bit1_def [unfolded Abs_bit1'_def]) apply (rule uminus_bit1_def [unfolded Abs_bit1'_def]) done instance bit0 :: (finite) comm_ring_1 by (rule bit0.comm_ring_1) instance bit1 :: (finite) comm_ring_1 by (rule bit1.comm_ring_1) interpretation bit0: mod_ring "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" .. interpretation bit1: mod_ring "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" .. text \Set up cases, induction, and arithmetic\ lemmas bit0_cases [case_names of_int, cases type: bit0] = bit0.cases lemmas bit1_cases [case_names of_int, cases type: bit1] = bit1.cases lemmas bit0_induct [case_names of_int, induct type: bit0] = bit0.induct lemmas bit1_induct [case_names of_int, induct type: bit1] = bit1.induct lemmas bit0_iszero_numeral [simp] = bit0.iszero_numeral lemmas bit1_iszero_numeral [simp] = bit1.iszero_numeral lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit0"] for dummy :: "'a::finite" lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit1"] for dummy :: "'a::finite" subsection \Order instances\ instantiation bit0 and bit1 :: (finite) linorder begin definition "a < b \ Rep_bit0 a < Rep_bit0 b" definition "a \ b \ Rep_bit0 a \ Rep_bit0 b" definition "a < b \ Rep_bit1 a < Rep_bit1 b" definition "a \ b \ Rep_bit1 a \ Rep_bit1 b" instance by(intro_classes) (auto simp add: less_eq_bit0_def less_bit0_def less_eq_bit1_def less_bit1_def Rep_bit0_inject Rep_bit1_inject) end lemma (in preorder) tranclp_less: "(<) \<^sup>+\<^sup>+ = (<)" by(auto simp add: fun_eq_iff intro: less_trans elim: tranclp.induct) instance bit0 and bit1 :: (finite) wellorder proof - have "wf {(x :: 'a bit0, y). x < y}" by(auto simp add: trancl_def tranclp_less intro!: finite_acyclic_wf acyclicI) thus "OFCLASS('a bit0, wellorder_class)" by(rule wf_wellorderI) intro_classes next have "wf {(x :: 'a bit1, y). x < y}" by(auto simp add: trancl_def tranclp_less intro!: finite_acyclic_wf acyclicI) thus "OFCLASS('a bit1, wellorder_class)" by(rule wf_wellorderI) intro_classes qed subsection \Code setup and type classes for code generation\ text \Code setup for \<^typ>\num0\ and \<^typ>\num1\\ definition Num0 :: num0 where "Num0 = Abs_num0 0" code_datatype Num0 instantiation num0 :: equal begin definition equal_num0 :: "num0 \ num0 \ bool" where "equal_num0 = (=)" instance by intro_classes (simp add: equal_num0_def) end lemma equal_num0_code [code]: "equal_class.equal Num0 Num0 = True" by(rule equal_refl) code_datatype "1 :: num1" instantiation num1 :: equal begin definition equal_num1 :: "num1 \ num1 \ bool" where "equal_num1 = (=)" instance by intro_classes (simp add: equal_num1_def) end lemma equal_num1_code [code]: "equal_class.equal (1 :: num1) 1 = True" by(rule equal_refl) instantiation num1 :: enum begin definition "enum_class.enum = [1 :: num1]" definition "enum_class.enum_all P = P (1 :: num1)" definition "enum_class.enum_ex P = P (1 :: num1)" instance by intro_classes (auto simp add: enum_num1_def enum_all_num1_def enum_ex_num1_def num1_eq_iff Ball_def) end instantiation num0 and num1 :: card_UNIV begin definition "finite_UNIV = Phantom(num0) False" definition "card_UNIV = Phantom(num0) 0" definition "finite_UNIV = Phantom(num1) True" definition "card_UNIV = Phantom(num1) 1" instance by intro_classes (simp_all add: finite_UNIV_num0_def card_UNIV_num0_def infinite_num0 finite_UNIV_num1_def card_UNIV_num1_def) end text \Code setup for \<^typ>\'a bit0\ and \<^typ>\'a bit1\\ declare bit0.Rep_inverse[code abstype] bit0.Rep_0[code abstract] bit0.Rep_1[code abstract] lemma Abs_bit0'_code [code abstract]: "Rep_bit0 (Abs_bit0' x :: 'a :: finite bit0) = x mod int (CARD('a bit0))" by(auto simp add: Abs_bit0'_def intro!: Abs_bit0_inverse) lemma inj_on_Abs_bit0: "inj_on (Abs_bit0 :: int \ 'a bit0) {0..<2 * int CARD('a :: finite)}" by(auto intro: inj_onI simp add: Abs_bit0_inject) declare bit1.Rep_inverse[code abstype] bit1.Rep_0[code abstract] bit1.Rep_1[code abstract] lemma Abs_bit1'_code [code abstract]: "Rep_bit1 (Abs_bit1' x :: 'a :: finite bit1) = x mod int (CARD('a bit1))" by(auto simp add: Abs_bit1'_def intro!: Abs_bit1_inverse) lemma inj_on_Abs_bit1: "inj_on (Abs_bit1 :: int \ 'a bit1) {0..<1 + 2 * int CARD('a :: finite)}" by(auto intro: inj_onI simp add: Abs_bit1_inject) instantiation bit0 and bit1 :: (finite) equal begin definition "equal_class.equal x y \ Rep_bit0 x = Rep_bit0 y" definition "equal_class.equal x y \ Rep_bit1 x = Rep_bit1 y" instance by intro_classes (simp_all add: equal_bit0_def equal_bit1_def Rep_bit0_inject Rep_bit1_inject) end instantiation bit0 :: (finite) enum begin definition "(enum_class.enum :: 'a bit0 list) = map (Abs_bit0' \ int) (upt 0 (CARD('a bit0)))" definition "enum_class.enum_all P = (\b :: 'a bit0 \ set enum_class.enum. P b)" definition "enum_class.enum_ex P = (\b :: 'a bit0 \ set enum_class.enum. P b)" instance proof show "distinct (enum_class.enum :: 'a bit0 list)" by (simp add: enum_bit0_def distinct_map inj_on_def Abs_bit0'_def Abs_bit0_inject) let ?Abs = "Abs_bit0 :: _ \ 'a bit0" interpret type_definition Rep_bit0 ?Abs "{0..<2 * int CARD('a)}" by (fact type_definition_bit0) have "UNIV = ?Abs ` {0..<2 * int CARD('a)}" by (simp add: Abs_image) also have "\ = ?Abs ` (int ` {0..<2 * CARD('a)})" by (simp add: image_int_atLeastLessThan) also have "\ = (?Abs \ int) ` {0..<2 * CARD('a)}" by (simp add: image_image cong: image_cong) also have "\ = set enum_class.enum" by (simp add: enum_bit0_def Abs_bit0'_def cong: image_cong_simp) finally show univ_eq: "(UNIV :: 'a bit0 set) = set enum_class.enum" . fix P :: "'a bit0 \ bool" show "enum_class.enum_all P = Ball UNIV P" and "enum_class.enum_ex P = Bex UNIV P" by(simp_all add: enum_all_bit0_def enum_ex_bit0_def univ_eq) qed end instantiation bit1 :: (finite) enum begin definition "(enum_class.enum :: 'a bit1 list) = map (Abs_bit1' \ int) (upt 0 (CARD('a bit1)))" definition "enum_class.enum_all P = (\b :: 'a bit1 \ set enum_class.enum. P b)" definition "enum_class.enum_ex P = (\b :: 'a bit1 \ set enum_class.enum. P b)" instance proof(intro_classes) show "distinct (enum_class.enum :: 'a bit1 list)" by(simp only: Abs_bit1'_def zmod_int[symmetric] enum_bit1_def distinct_map Suc_eq_plus1 card_bit1 o_apply inj_on_def) (clarsimp simp add: Abs_bit1_inject) let ?Abs = "Abs_bit1 :: _ \ 'a bit1" interpret type_definition Rep_bit1 ?Abs "{0..<1 + 2 * int CARD('a)}" by (fact type_definition_bit1) have "UNIV = ?Abs ` {0..<1 + 2 * int CARD('a)}" by (simp add: Abs_image) also have "\ = ?Abs ` (int ` {0..<1 + 2 * CARD('a)})" by (simp add: image_int_atLeastLessThan) also have "\ = (?Abs \ int) ` {0..<1 + 2 * CARD('a)}" by (simp add: image_image cong: image_cong) finally show univ_eq: "(UNIV :: 'a bit1 set) = set enum_class.enum" by (simp only: enum_bit1_def set_map set_upt) (simp add: Abs_bit1'_def cong: image_cong_simp) fix P :: "'a bit1 \ bool" show "enum_class.enum_all P = Ball UNIV P" and "enum_class.enum_ex P = Bex UNIV P" by(simp_all add: enum_all_bit1_def enum_ex_bit1_def univ_eq) qed end instantiation bit0 and bit1 :: (finite) finite_UNIV begin definition "finite_UNIV = Phantom('a bit0) True" definition "finite_UNIV = Phantom('a bit1) True" instance by intro_classes (simp_all add: finite_UNIV_bit0_def finite_UNIV_bit1_def) end instantiation bit0 and bit1 :: ("{finite,card_UNIV}") card_UNIV begin definition "card_UNIV = Phantom('a bit0) (2 * of_phantom (card_UNIV :: 'a card_UNIV))" definition "card_UNIV = Phantom('a bit1) (1 + 2 * of_phantom (card_UNIV :: 'a card_UNIV))" instance by intro_classes (simp_all add: card_UNIV_bit0_def card_UNIV_bit1_def card_UNIV) end subsection \Syntax\ syntax "_NumeralType" :: "num_token => type" ("_") "_NumeralType0" :: type ("0") "_NumeralType1" :: type ("1") translations (type) "1" == (type) "num1" (type) "0" == (type) "num0" parse_translation \ let fun mk_bintype n = let fun mk_bit 0 = Syntax.const \<^type_syntax>\bit0\ | mk_bit 1 = Syntax.const \<^type_syntax>\bit1\; fun bin_of n = if n = 1 then Syntax.const \<^type_syntax>\num1\ else if n = 0 then Syntax.const \<^type_syntax>\num0\ else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r $ bin_of q end; in bin_of n end; fun numeral_tr [Free (str, _)] = mk_bintype (the (Int.fromString str)) | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_NumeralType\, K numeral_tr)] end \ print_translation \ let fun int_of [] = 0 | int_of (b :: bs) = b + 2 * int_of bs; fun bin_of (Const (\<^type_syntax>\num0\, _)) = [] | bin_of (Const (\<^type_syntax>\num1\, _)) = [1] | bin_of (Const (\<^type_syntax>\bit0\, _) $ bs) = 0 :: bin_of bs | bin_of (Const (\<^type_syntax>\bit1\, _) $ bs) = 1 :: bin_of bs | bin_of t = raise TERM ("bin_of", [t]); fun bit_tr' b [t] = let val rev_digs = b :: bin_of t handle TERM _ => raise Match val i = int_of rev_digs; val num = string_of_int (abs i); in Syntax.const \<^syntax_const>\_NumeralType\ $ Syntax.free num end | bit_tr' b _ = raise Match; in [(\<^type_syntax>\bit0\, K (bit_tr' 0)), (\<^type_syntax>\bit1\, K (bit_tr' 1))] end \ subsection \Examples\ lemma "CARD(0) = 0" by simp lemma "CARD(17) = 17" by simp lemma "CHAR(23) = 23" by simp lemma "8 * 11 ^ 3 - 6 = (2::5)" by simp end diff --git a/src/HOL/Library/Omega_Words_Fun.thy b/src/HOL/Library/Omega_Words_Fun.thy --- a/src/HOL/Library/Omega_Words_Fun.thy +++ b/src/HOL/Library/Omega_Words_Fun.thy @@ -1,967 +1,967 @@ (* Author: Stefan Merz Author: Salomon Sickert Author: Julian Brunner Author: Peter Lammich *) section \$\omega$-words\ theory Omega_Words_Fun imports Infinite_Set begin text \Note: This theory is based on Stefan Merz's work.\ text \ Automata recognize languages, which are sets of words. For the theory of $\omega$-automata, we are mostly interested in $\omega$-words, but it is sometimes useful to reason about finite words, too. We are modeling finite words as lists; this lets us benefit from the existing library. Other formalizations could be investigated, such as representing words as functions whose domains are initial intervals of the natural numbers. \ subsection \Type declaration and elementary operations\ text \ We represent $\omega$-words as functions from the natural numbers to the alphabet type. Other possible formalizations include a coinductive definition or a uniform encoding of finite and infinite words, as studied by Müller et al. \ type_synonym 'a word = "nat \ 'a" text \ We can prefix a finite word to an $\omega$-word, and a way to obtain an $\omega$-word from a finite, non-empty word is by $\omega$-iteration. \ definition conc :: "['a list, 'a word] \ 'a word" (infixr \\\ 65) where "w \ x == \n. if n < length w then w!n else x (n - length w)" definition iter :: "'a list \ 'a word" (\(_\<^sup>\)\ [1000]) where "iter w == if w = [] then undefined else (\n. w!(n mod (length w)))" lemma conc_empty[simp]: "[] \ w = w" unfolding conc_def by auto lemma conc_fst[simp]: "n < length w \ (w \ x) n = w!n" by (simp add: conc_def) lemma conc_snd[simp]: "\(n < length w) \ (w \ x) n = x (n - length w)" by (simp add: conc_def) lemma iter_nth [simp]: "0 < length w \ w\<^sup>\ n = w!(n mod (length w))" by (simp add: iter_def) lemma conc_conc[simp]: "u \ v \ w = (u @ v) \ w" (is "?lhs = ?rhs") proof fix n have u: "n < length u \ ?lhs n = ?rhs n" by (simp add: conc_def nth_append) have v: "\ \(n < length u); n < length u + length v \ \ ?lhs n = ?rhs n" by (simp add: conc_def nth_append, arith) have w: "\(n < length u + length v) \ ?lhs n = ?rhs n" by (simp add: conc_def nth_append, arith) from u v w show "?lhs n = ?rhs n" by blast qed lemma range_conc[simp]: "range (w\<^sub>1 \ w\<^sub>2) = set w\<^sub>1 \ range w\<^sub>2" proof (intro equalityI subsetI) fix a assume "a \ range (w\<^sub>1 \ w\<^sub>2)" then obtain i where 1: "a = (w\<^sub>1 \ w\<^sub>2) i" by auto then show "a \ set w\<^sub>1 \ range w\<^sub>2" unfolding 1 by (cases "i < length w\<^sub>1") simp_all next fix a assume a: "a \ set w\<^sub>1 \ range w\<^sub>2" then show "a \ range (w\<^sub>1 \ w\<^sub>2)" proof assume "a \ set w\<^sub>1" then obtain i where 1: "i < length w\<^sub>1" "a = w\<^sub>1 ! i" using in_set_conv_nth by metis show ?thesis proof show "a = (w\<^sub>1 \ w\<^sub>2) i" using 1 by auto show "i \ UNIV" by rule qed next assume "a \ range w\<^sub>2" then obtain i where 1: "a = w\<^sub>2 i" by auto show ?thesis proof show "a = (w\<^sub>1 \ w\<^sub>2) (length w\<^sub>1 + i)" using 1 by simp show "length w\<^sub>1 + i \ UNIV" by rule qed qed qed lemma iter_unroll: "0 < length w \ w\<^sup>\ = w \ w\<^sup>\" - by (rule ext) (simp add: conc_def mod_geq) + by (simp add: fun_eq_iff mod_if) subsection \Subsequence, Prefix, and Suffix\ definition suffix :: "[nat, 'a word] \ 'a word" where "suffix k x \ \n. x (k+n)" definition subsequence :: "'a word \ nat \ nat \ 'a list" (\_ [_ \ _]\ 900) where "subsequence w i j \ map w [i.. 'a word \ 'a list" where "prefix n w \ subsequence w 0 n" lemma suffix_nth [simp]: "(suffix k x) n = x (k+n)" by (simp add: suffix_def) lemma suffix_0 [simp]: "suffix 0 x = x" by (simp add: suffix_def) lemma suffix_suffix [simp]: "suffix m (suffix k x) = suffix (k+m) x" by (rule ext) (simp add: suffix_def add.assoc) lemma subsequence_append: "prefix (i + j) w = prefix i w @ (w [i \ i + j])" unfolding map_append[symmetric] upt_add_eq_append[OF le0] subsequence_def .. lemma subsequence_drop[simp]: "drop i (w [j \ k]) = w [j + i \ k]" by (simp add: subsequence_def drop_map) lemma subsequence_empty[simp]: "w [i \ j] = [] \ j \ i" by (auto simp add: subsequence_def) lemma subsequence_length[simp]: "length (subsequence w i j) = j - i" by (simp add: subsequence_def) lemma subsequence_nth[simp]: "k < j - i \ (w [i \ j]) ! k = w (i + k)" unfolding subsequence_def by auto lemma subseq_to_zero[simp]: "w[i\0] = []" by simp lemma subseq_to_smaller[simp]: "i\j \ w[i\j] = []" by simp lemma subseq_to_Suc[simp]: "i\j \ w [i \ Suc j] = w [ i \ j ] @ [w j]" by (auto simp: subsequence_def) lemma subsequence_singleton[simp]: "w [i \ Suc i] = [w i]" by (auto simp: subsequence_def) lemma subsequence_prefix_suffix: "prefix (j - i) (suffix i w) = w [i \ j]" proof (cases "i \ j") case True have "w [i \ j] = map w (map (\n. n + i) [0.. = map (\n. w (n + i)) [0.. (suffix n x)" by (rule ext) (simp add: subsequence_def conc_def) declare prefix_suffix[symmetric, simp] lemma word_split: obtains v\<^sub>1 v\<^sub>2 where "v = v\<^sub>1 \ v\<^sub>2" "length v\<^sub>1 = k" proof show "v = prefix k v \ suffix k v" by (rule prefix_suffix) show "length (prefix k v) = k" by simp qed lemma set_subsequence[simp]: "set (w[i\j]) = w`{i.. k]) = w [j \ min (j + i) k]" by (simp add: subsequence_def take_map min_def) lemma subsequence_shift[simp]: "(suffix i w) [j \ k] = w [i + j \ i + k]" by (metis add_diff_cancel_left subsequence_prefix_suffix suffix_suffix) lemma suffix_subseq_join[simp]: "i \ j \ v [i \ j] \ suffix j v = suffix i v" by (metis (no_types, lifting) Nat.add_0_right le_add_diff_inverse prefix_suffix subsequence_shift suffix_suffix) lemma prefix_conc_fst[simp]: assumes "j \ length w" shows "prefix j (w \ w') = take j w" proof - have "\i < j. (prefix j (w \ w')) ! i = (take j w) ! i" using assms by (simp add: conc_fst subsequence_def) thus ?thesis by (simp add: assms list_eq_iff_nth_eq min.absorb2) qed lemma prefix_conc_snd[simp]: assumes "n \ length u" shows "prefix n (u \ v) = u @ prefix (n - length u) v" proof (intro nth_equalityI) show "length (prefix n (u \ v)) = length (u @ prefix (n - length u) v)" using assms by simp fix i assume "i < length (prefix n (u \ v))" then show "prefix n (u \ v) ! i = (u @ prefix (n - length u) v) ! i" by (cases "i < length u") (auto simp: nth_append) qed lemma prefix_conc_length[simp]: "prefix (length w) (w \ w') = w" by simp lemma suffix_conc_fst[simp]: assumes "n \ length u" shows "suffix n (u \ v) = drop n u \ v" proof show "suffix n (u \ v) i = (drop n u \ v) i" for i using assms by (cases "n + i < length u") (auto simp: algebra_simps) qed lemma suffix_conc_snd[simp]: assumes "n \ length u" shows "suffix n (u \ v) = suffix (n - length u) v" proof show "suffix n (u \ v) i = suffix (n - length u) v i" for i using assms by simp qed lemma suffix_conc_length[simp]: "suffix (length w) (w \ w') = w'" unfolding conc_def by force lemma concat_eq[iff]: assumes "length v\<^sub>1 = length v\<^sub>2" shows "v\<^sub>1 \ u\<^sub>1 = v\<^sub>2 \ u\<^sub>2 \ v\<^sub>1 = v\<^sub>2 \ u\<^sub>1 = u\<^sub>2" (is "?lhs \ ?rhs") proof assume ?lhs then have 1: "(v\<^sub>1 \ u\<^sub>1) i = (v\<^sub>2 \ u\<^sub>2) i" for i by auto show ?rhs proof (intro conjI ext nth_equalityI) show "length v\<^sub>1 = length v\<^sub>2" by (rule assms(1)) next fix i assume 2: "i < length v\<^sub>1" have 3: "i < length v\<^sub>2" using assms(1) 2 by simp show "v\<^sub>1 ! i = v\<^sub>2 ! i" using 1[of i] 2 3 by simp next show "u\<^sub>1 i = u\<^sub>2 i" for i using 1[of "length v\<^sub>1 + i"] assms(1) by simp qed next assume ?rhs then show ?lhs by simp qed lemma same_concat_eq[iff]: "u \ v = u \ w \ v = w" by simp lemma comp_concat[simp]: "f \ u \ v = map f u \ (f \ v)" proof fix i show "(f \ u \ v) i = (map f u \ (f \ v)) i" by (cases "i < length u") simp_all qed subsection \Prepending\ primrec build :: "'a \ 'a word \ 'a word" (infixr \##\ 65) where "(a ## w) 0 = a" | "(a ## w) (Suc i) = w i" lemma build_eq[iff]: "a\<^sub>1 ## w\<^sub>1 = a\<^sub>2 ## w\<^sub>2 \ a\<^sub>1 = a\<^sub>2 \ w\<^sub>1 = w\<^sub>2" proof assume 1: "a\<^sub>1 ## w\<^sub>1 = a\<^sub>2 ## w\<^sub>2" have 2: "(a\<^sub>1 ## w\<^sub>1) i = (a\<^sub>2 ## w\<^sub>2) i" for i using 1 by auto show "a\<^sub>1 = a\<^sub>2 \ w\<^sub>1 = w\<^sub>2" proof (intro conjI ext) show "a\<^sub>1 = a\<^sub>2" using 2[of "0"] by simp show "w\<^sub>1 i = w\<^sub>2 i" for i using 2[of "Suc i"] by simp qed next assume 1: "a\<^sub>1 = a\<^sub>2 \ w\<^sub>1 = w\<^sub>2" show "a\<^sub>1 ## w\<^sub>1 = a\<^sub>2 ## w\<^sub>2" using 1 by simp qed lemma build_cons[simp]: "(a # u) \ v = a ## u \ v" proof fix i show "((a # u) \ v) i = (a ## u \ v) i" proof (cases i) case 0 show ?thesis unfolding 0 by simp next case (Suc j) show ?thesis unfolding Suc by (cases "j < length u", simp+) qed qed lemma build_append[simp]: "(w @ a # u) \ v = w \ a ## u \ v" unfolding conc_conc[symmetric] by simp lemma build_first[simp]: "w 0 ## suffix (Suc 0) w = w" proof show "(w 0 ## suffix (Suc 0) w) i = w i" for i by (cases i) simp_all qed lemma build_split[intro]: "w = w 0 ## suffix 1 w" by simp lemma build_range[simp]: "range (a ## w) = insert a (range w)" proof safe show "(a ## w) i \ range w \ (a ## w) i = a" for i by (cases i) auto show "a \ range (a ## w)" proof (rule range_eqI) show "a = (a ## w) 0" by simp qed show "w i \ range (a ## w)" for i proof (rule range_eqI) show "w i = (a ## w) (Suc i)" by simp qed qed lemma suffix_singleton_suffix[simp]: "w i ## suffix (Suc i) w = suffix i w" using suffix_subseq_join[of i "Suc i" w] by simp text \Find the first occurrence of a letter from a given set\ lemma word_first_split_set: assumes "A \ range w \ {}" obtains u a v where "w = u \ [a] \ v" "A \ set u = {}" "a \ A" proof - define i where "i = (LEAST i. w i \ A)" show ?thesis proof show "w = prefix i w \ [w i] \ suffix (Suc i) w" by simp show "A \ set (prefix i w) = {}" apply safe subgoal premises prems for a proof - from prems obtain k where 3: "k < i" "w k = a" by auto have 4: "w k \ A" using not_less_Least 3(1) unfolding i_def . show ?thesis using prems(1) 3(2) 4 by auto qed done show "w i \ A" using LeastI assms(1) unfolding i_def by fast qed qed subsection \The limit set of an $\omega$-word\ text \ The limit set (also called infinity set) of an $\omega$-word is the set of letters that appear infinitely often in the word. This set plays an important role in defining acceptance conditions of $\omega$-automata. \ definition limit :: "'a word \ 'a set" where "limit x \ {a . \\<^sub>\n . x n = a}" lemma limit_iff_frequent: "a \ limit x \ (\\<^sub>\n . x n = a)" by (simp add: limit_def) text \ The following is a different way to define the limit, using the reverse image, making the laws about reverse image applicable to the limit set. (Might want to change the definition above?) \ lemma limit_vimage: "(a \ limit x) = infinite (x -` {a})" by (simp add: limit_def Inf_many_def vimage_def) lemma two_in_limit_iff: "({a, b} \ limit x) = ((\n. x n =a ) \ (\n. x n = a \ (\m>n. x m = b)) \ (\m. x m = b \ (\n>m. x n = a)))" (is "?lhs = (?r1 \ ?r2 \ ?r3)") proof assume lhs: "?lhs" hence 1: "?r1" by (auto simp: limit_def elim: INFM_EX) from lhs have "\n. \m>n. x m = b" by (auto simp: limit_def INFM_nat) hence 2: "?r2" by simp from lhs have "\m. \n>m. x n = a" by (auto simp: limit_def INFM_nat) hence 3: "?r3" by simp from 1 2 3 show "?r1 \ ?r2 \ ?r3" by simp next assume "?r1 \ ?r2 \ ?r3" hence 1: "?r1" and 2: "?r2" and 3: "?r3" by simp+ have infa: "\m. \n\m. x n = a" proof fix m show "\n\m. x n = a" (is "?A m") proof (induct m) from 1 show "?A 0" by simp next fix m assume ih: "?A m" then obtain n where n: "n \ m" "x n = a" by auto with 2 obtain k where k: "k>n" "x k = b" by auto with 3 obtain l where l: "l>k" "x l = a" by auto from n k l have "l \ Suc m" by auto with l show "?A (Suc m)" by auto qed qed hence infa': "\\<^sub>\n. x n = a" by (simp add: INFM_nat_le) have "\n. \m>n. x m = b" proof fix n from infa obtain k where k1: "k\n" and k2: "x k = a" by auto from 2 k2 obtain l where l1: "l>k" and l2: "x l = b" by auto from k1 l1 have "l > n" by auto with l2 show "\m>n. x m = b" by auto qed hence "\\<^sub>\m. x m = b" by (simp add: INFM_nat) with infa' show "?lhs" by (auto simp: limit_def) qed text \ For $\omega$-words over a finite alphabet, the limit set is non-empty. Moreover, from some position onward, any such word contains only letters from its limit set. \ lemma limit_nonempty: assumes fin: "finite (range x)" shows "\a. a \ limit x" proof - from fin obtain a where "a \ range x \ infinite (x -` {a})" by (rule inf_img_fin_domE) auto hence "a \ limit x" by (auto simp add: limit_vimage) thus ?thesis .. qed lemmas limit_nonemptyE = limit_nonempty[THEN exE] lemma limit_inter_INF: assumes hyp: "limit w \ S \ {}" shows "\\<^sub>\ n. w n \ S" proof - from hyp obtain x where "\\<^sub>\ n. w n = x" and "x \ S" by (auto simp add: limit_def) thus ?thesis by (auto elim: INFM_mono) qed text \ The reverse implication is true only if $S$ is finite. \ lemma INF_limit_inter: assumes hyp: "\\<^sub>\ n. w n \ S" and fin: "finite (S \ range w)" shows "\a. a \ limit w \ S" proof (rule ccontr) assume contra: "\(\a. a \ limit w \ S)" hence "\a\S. finite {n. w n = a}" by (auto simp add: limit_def Inf_many_def) with fin have "finite (UN a:S \ range w. {n. w n = a})" by auto moreover have "(UN a:S \ range w. {n. w n = a}) = {n. w n \ S}" by auto moreover note hyp ultimately show "False" by (simp add: Inf_many_def) qed lemma fin_ex_inf_eq_limit: "finite A \ (\\<^sub>\i. w i \ A) \ limit w \ A \ {}" by (metis INF_limit_inter equals0D finite_Int limit_inter_INF) lemma limit_in_range_suffix: "limit x \ range (suffix k x)" proof fix a assume "a \ limit x" then obtain l where kl: "k < l" and xl: "x l = a" by (auto simp add: limit_def INFM_nat) from kl obtain m where "l = k+m" by (auto simp add: less_iff_Suc_add) with xl show "a \ range (suffix k x)" by auto qed lemma limit_in_range: "limit r \ range r" using limit_in_range_suffix[of r 0] by simp lemmas limit_in_range_suffixD = limit_in_range_suffix[THEN subsetD] lemma limit_subset: "limit f \ f ` {n..}" using limit_in_range_suffix[of f n] unfolding suffix_def by auto theorem limit_is_suffix: assumes fin: "finite (range x)" shows "\k. limit x = range (suffix k x)" proof - have "\k. range (suffix k x) \ limit x" proof - \ \The set of letters that are not in the limit is certainly finite.\ from fin have "finite (range x - limit x)" by simp \ \Moreover, any such letter occurs only finitely often\ moreover have "\a \ range x - limit x. finite (x -` {a})" by (auto simp add: limit_vimage) \ \Thus, there are only finitely many occurrences of such letters.\ ultimately have "finite (UN a : range x - limit x. x -` {a})" by (blast intro: finite_UN_I) \ \Therefore these occurrences are within some initial interval.\ then obtain k where "(UN a : range x - limit x. x -` {a}) \ {.. \This is just the bound we are looking for.\ hence "\m. k \ m \ x m \ limit x" by (auto simp add: limit_vimage) hence "range (suffix k x) \ limit x" by auto thus ?thesis .. qed then obtain k where "range (suffix k x) \ limit x" .. with limit_in_range_suffix have "limit x = range (suffix k x)" by (rule subset_antisym) thus ?thesis .. qed lemmas limit_is_suffixE = limit_is_suffix[THEN exE] text \ The limit set enjoys some simple algebraic laws with respect to concatenation, suffixes, iteration, and renaming. \ theorem limit_conc [simp]: "limit (w \ x) = limit x" proof (auto) fix a assume a: "a \ limit (w \ x)" have "\m. \n. m x n = a" proof fix m from a obtain n where "m + length w < n \ (w \ x) n = a" by (auto simp add: limit_def Inf_many_def infinite_nat_iff_unbounded) hence "m < n - length w \ x (n - length w) = a" by (auto simp add: conc_def) thus "\n. m x n = a" .. qed hence "infinite {n . x n = a}" by (simp add: infinite_nat_iff_unbounded) thus "a \ limit x" by (simp add: limit_def Inf_many_def) next fix a assume a: "a \ limit x" have "\m. length w < m \ (\n. m (w \ x) n = a)" proof (clarify) fix m assume m: "length w < m" with a obtain n where "m - length w < n \ x n = a" by (auto simp add: limit_def Inf_many_def infinite_nat_iff_unbounded) with m have "m < n + length w \ (w \ x) (n + length w) = a" by (simp add: conc_def, arith) thus "\n. m (w \ x) n = a" .. qed hence "infinite {n . (w \ x) n = a}" by (simp add: unbounded_k_infinite) thus "a \ limit (w \ x)" by (simp add: limit_def Inf_many_def) qed theorem limit_suffix [simp]: "limit (suffix n x) = limit x" proof - have "x = (prefix n x) \ (suffix n x)" by (simp add: prefix_suffix) hence "limit x = limit (prefix n x \ suffix n x)" by simp also have "\ = limit (suffix n x)" by (rule limit_conc) finally show ?thesis by (rule sym) qed theorem limit_iter [simp]: assumes nempty: "0 < length w" shows "limit w\<^sup>\ = set w" proof have "limit w\<^sup>\ \ range w\<^sup>\" by (auto simp add: limit_def dest: INFM_EX) also from nempty have "\ \ set w" by auto finally show "limit w\<^sup>\ \ set w" . next { fix a assume a: "a \ set w" then obtain k where k: "k < length w \ w!k = a" by (auto simp add: set_conv_nth) \ \the following bound is terrible, but it simplifies the proof\ from nempty k have "\m. w\<^sup>\ ((Suc m)*(length w) + k) = a" by (simp add: mod_add_left_eq [symmetric]) moreover \ \why is the following so hard to prove??\ have "\m. m < (Suc m)*(length w) + k" proof fix m from nempty have "1 \ length w" by arith hence "m*1 \ m*length w" by simp hence "m \ m*length w" by simp with nempty have "m < length w + (m*length w) + k" by arith thus "m < (Suc m)*(length w) + k" by simp qed moreover note nempty ultimately have "a \ limit w\<^sup>\" by (auto simp add: limit_iff_frequent INFM_nat) } then show "set w \ limit w\<^sup>\" by auto qed lemma limit_o [simp]: assumes a: "a \ limit w" shows "f a \ limit (f \ w)" proof - from a have "\\<^sub>\n. w n = a" by (simp add: limit_iff_frequent) hence "\\<^sub>\n. f (w n) = f a" by (rule INFM_mono, simp) thus "f a \ limit (f \ w)" by (simp add: limit_iff_frequent) qed text \ The converse relation is not true in general: $f(a)$ can be in the limit of $f \circ w$ even though $a$ is not in the limit of $w$. However, \limit\ commutes with renaming if the function is injective. More generally, if $f(a)$ is the image of only finitely many elements, some of these must be in the limit of $w$. \ lemma limit_o_inv: assumes fin: "finite (f -` {x})" and x: "x \ limit (f \ w)" shows "\a \ (f -` {x}). a \ limit w" proof (rule ccontr) assume contra: "\ ?thesis" \ \hence, every element in the pre-image occurs only finitely often\ then have "\a \ (f -` {x}). finite {n. w n = a}" by (simp add: limit_def Inf_many_def) \ \so there are only finitely many occurrences of any such element\ with fin have "finite (\ a \ (f -` {x}). {n. w n = a})" by auto \ \these are precisely those positions where $x$ occurs in $f \circ w$\ moreover have "(\ a \ (f -` {x}). {n. w n = a}) = {n. f(w n) = x}" by auto ultimately \ \so $x$ can occur only finitely often in the translated word\ have "finite {n. f(w n) = x}" by simp \ \\ldots\ which yields a contradiction\ with x show "False" by (simp add: limit_def Inf_many_def) qed theorem limit_inj [simp]: assumes inj: "inj f" shows "limit (f \ w) = f ` (limit w)" proof show "f ` limit w \ limit (f \ w)" by auto show "limit (f \ w) \ f ` limit w" proof fix x assume x: "x \ limit (f \ w)" from inj have "finite (f -` {x})" by (blast intro: finite_vimageI) with x obtain a where a: "a \ (f -` {x}) \ a \ limit w" by (blast dest: limit_o_inv) thus "x \ f ` (limit w)" by auto qed qed lemma limit_inter_empty: assumes fin: "finite (range w)" assumes hyp: "limit w \ S = {}" shows "\\<^sub>\n. w n \ S" proof - from fin obtain k where k_def: "limit w = range (suffix k w)" using limit_is_suffix by blast have "w (k + k') \ S" for k' using hyp unfolding k_def suffix_def image_def by blast thus ?thesis unfolding MOST_nat_le using le_Suc_ex by blast qed text \If the limit is the suffix of the sequence's range, we may increase the suffix index arbitrarily\ lemma limit_range_suffix_incr: assumes "limit r = range (suffix i r)" assumes "j\i" shows "limit r = range (suffix j r)" (is "?lhs = ?rhs") proof - have "?lhs = range (suffix i r)" using assms by simp moreover have "\ \ ?rhs" using \j\i\ by (metis (mono_tags, lifting) assms(2) image_subsetI le_Suc_ex range_eqI suffix_def suffix_suffix) moreover have "\ \ ?lhs" by (rule limit_in_range_suffix) ultimately show "?lhs = ?rhs" by (metis antisym_conv limit_in_range_suffix) qed text \For two finite sequences, we can find a common suffix index such that the limits can be represented as these suffixes' ranges.\ lemma common_range_limit: assumes "finite (range x)" and "finite (range y)" obtains i where "limit x = range (suffix i x)" and "limit y = range (suffix i y)" proof - obtain i j where 1: "limit x = range (suffix i x)" and 2: "limit y = range (suffix j y)" using assms limit_is_suffix by metis have "limit x = range (suffix (max i j) x)" and "limit y = range (suffix (max i j) y)" using limit_range_suffix_incr[OF 1] limit_range_suffix_incr[OF 2] by auto thus ?thesis using that by metis qed subsection \Index sequences and piecewise definitions\ text \ A word can be defined piecewise: given a sequence of words $w_0, w_1, \ldots$ and a strictly increasing sequence of integers $i_0, i_1, \ldots$ where $i_0=0$, a single word is obtained by concatenating subwords of the $w_n$ as given by the integers: the resulting word is \[ (w_0)_{i_0} \ldots (w_0)_{i_1-1} (w_1)_{i_1} \ldots (w_1)_{i_2-1} \ldots \] We prepare the field by proving some trivial facts about such sequences of indexes. \ definition idx_sequence :: "nat word \ bool" where "idx_sequence idx \ (idx 0 = 0) \ (\n. idx n < idx (Suc n))" lemma idx_sequence_less: assumes iseq: "idx_sequence idx" shows "idx n < idx (Suc(n+k))" proof (induct k) from iseq show "idx n < idx (Suc (n + 0))" by (simp add: idx_sequence_def) next fix k assume ih: "idx n < idx (Suc(n+k))" from iseq have "idx (Suc(n+k)) < idx (Suc(n + Suc k))" by (simp add: idx_sequence_def) with ih show "idx n < idx (Suc(n + Suc k))" by (rule less_trans) qed lemma idx_sequence_inj: assumes iseq: "idx_sequence idx" and eq: "idx m = idx n" shows "m = n" proof (cases m n rule: linorder_cases) case greater then obtain k where "m = Suc(n+k)" by (auto simp add: less_iff_Suc_add) with iseq have "idx n < idx m" by (simp add: idx_sequence_less) with eq show ?thesis by simp next case less then obtain k where "n = Suc(m+k)" by (auto simp add: less_iff_Suc_add) with iseq have "idx m < idx n" by (simp add: idx_sequence_less) with eq show ?thesis by simp qed lemma idx_sequence_mono: assumes iseq: "idx_sequence idx" and m: "m \ n" shows "idx m \ idx n" proof (cases "m=n") case True thus ?thesis by simp next case False with m have "m < n" by simp then obtain k where "n = Suc(m+k)" by (auto simp add: less_iff_Suc_add) with iseq have "idx m < idx n" by (simp add: idx_sequence_less) thus ?thesis by simp qed text \ Given an index sequence, every natural number is contained in the interval defined by two adjacent indexes, and in fact this interval is determined uniquely. \ lemma idx_sequence_idx: assumes "idx_sequence idx" shows "idx k \ {idx k ..< idx (Suc k)}" using assms by (auto simp add: idx_sequence_def) lemma idx_sequence_interval: assumes iseq: "idx_sequence idx" shows "\k. n \ {idx k ..< idx (Suc k) }" (is "?P n" is "\k. ?in n k") proof (induct n) from iseq have "0 = idx 0" by (simp add: idx_sequence_def) moreover from iseq have "idx 0 \ {idx 0 ..< idx (Suc 0) }" by (rule idx_sequence_idx) ultimately show "?P 0" by auto next fix n assume "?P n" then obtain k where k: "?in n k" .. show "?P (Suc n)" proof (cases "Suc n < idx (Suc k)") case True with k have "?in (Suc n) k" by simp thus ?thesis .. next case False with k have "Suc n = idx (Suc k)" by auto with iseq have "?in (Suc n) (Suc k)" by (simp add: idx_sequence_def) thus ?thesis .. qed qed lemma idx_sequence_interval_unique: assumes iseq: "idx_sequence idx" and k: "n \ {idx k ..< idx (Suc k)}" and m: "n \ {idx m ..< idx (Suc m)}" shows "k = m" proof (cases k m rule: linorder_cases) case less hence "Suc k \ m" by simp with iseq have "idx (Suc k) \ idx m" by (rule idx_sequence_mono) with m have "idx (Suc k) \ n" by auto with k have "False" by simp thus ?thesis .. next case greater hence "Suc m \ k" by simp with iseq have "idx (Suc m) \ idx k" by (rule idx_sequence_mono) with k have "idx (Suc m) \ n" by auto with m have "False" by simp thus ?thesis .. qed lemma idx_sequence_unique_interval: assumes iseq: "idx_sequence idx" shows "\! k. n \ {idx k ..< idx (Suc k) }" proof (rule ex_ex1I) from iseq show "\k. n \ {idx k ..< idx (Suc k)}" by (rule idx_sequence_interval) next fix k y assume "n \ {idx k.. {idx y.. Now we can define the piecewise construction of a word using an index sequence. \ definition merge :: "'a word word \ nat word \ 'a word" where "merge ws idx \ \n. let i = THE i. n \ {idx i ..< idx (Suc i) } in ws i n" lemma merge: assumes idx: "idx_sequence idx" and n: "n \ {idx i ..< idx (Suc i)}" shows "merge ws idx n = ws i n" proof - from n have "(THE k. n \ {idx k ..< idx (Suc k) }) = i" by (rule the_equality[OF _ sym[OF idx_sequence_interval_unique[OF idx n]]]) simp thus ?thesis by (simp add: merge_def Let_def) qed lemma merge0: assumes idx: "idx_sequence idx" shows "merge ws idx 0 = ws 0 0" proof (rule merge[OF idx]) from idx have "idx 0 < idx (Suc 0)" unfolding idx_sequence_def by blast with idx show "0 \ {idx 0 ..< idx (Suc 0)}" by (simp add: idx_sequence_def) qed lemma merge_Suc: assumes idx: "idx_sequence idx" and n: "n \ {idx i ..< idx (Suc i)}" shows "merge ws idx (Suc n) = (if Suc n = idx (Suc i) then ws (Suc i) else ws i) (Suc n)" proof auto assume eq: "Suc n = idx (Suc i)" from idx have "idx (Suc i) < idx (Suc(Suc i))" unfolding idx_sequence_def by blast with eq idx show "merge ws idx (idx (Suc i)) = ws (Suc i) (idx (Suc i))" by (simp add: merge) next assume neq: "Suc n \ idx (Suc i)" with n have "Suc n \ {idx i ..< idx (Suc i) }" by auto with idx show "merge ws idx (Suc n) = ws i (Suc n)" by (rule merge) qed end diff --git a/src/HOL/Library/Word.thy b/src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy +++ b/src/HOL/Library/Word.thy @@ -1,4506 +1,4495 @@ (* Title: HOL/Library/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA, et. al. *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" begin subsection \Preliminaries\ lemma signed_take_bit_decr_length_iff: \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by (cases \LENGTH('a)\) (simp_all add: signed_take_bit_eq_iff_take_bit_eq) subsection \Fundamentals\ subsubsection \Type definition\ quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) hide_const (open) rep \ \only for foundational purpose\ hide_const (open) Word \ \only for code generation\ subsubsection \Basic arithmetic\ instantiation word :: (len) comm_ring_1 begin lift_definition zero_word :: \'a word\ is 0 . lift_definition one_word :: \'a word\ is 1 . lift_definition plus_word :: \'a word \ 'a word \ 'a word\ is \(+)\ by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition minus_word :: \'a word \ 'a word \ 'a word\ is \(-)\ by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lift_definition uminus_word :: \'a word \ 'a word\ is uminus by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) lift_definition times_word :: \'a word \ 'a word \ 'a word\ is \(*)\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by (standard; transfer) (simp_all add: algebra_simps) end context includes lifting_syntax notes power_transfer [transfer_rule] transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] transfer_rule_of_nat [transfer_rule] transfer_rule_of_int [transfer_rule] begin lemma power_transfer_word [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) of_bool of_bool\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) numeral numeral\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) int of_nat\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) (\k. k) of_int\ proof - have \((=) ===> pcr_word) of_int of_int\ by transfer_prover then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed end lemma exp_eq_zero_iff [simp]: \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ by transfer auto lemma word_exp_length_eq_0 [simp]: \(2 :: 'a::len word) ^ LENGTH('a) = 0\ by simp subsubsection \Basic tool setup\ ML_file \Tools/word_lib.ML\ subsubsection \Basic code generation setup\ context begin qualified lift_definition the_int :: \'a::len word \ int\ is \take_bit LENGTH('a)\ . end lemma [code abstype]: \Word.Word (Word.the_int w) = w\ by transfer simp lemma Word_eq_word_of_int [code_post, simp]: \Word.Word = of_int\ by (rule; transfer) simp quickcheck_generator word constructors: \0 :: 'a::len word\, \numeral :: num \ 'a::len word\ instantiation word :: (len) equal begin lift_definition equal_word :: \'a word \ 'a word \ bool\ is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by simp instance by (standard; transfer) rule end lemma [code]: \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ by transfer (simp add: equal) lemma [code]: \Word.the_int 0 = 0\ by transfer simp lemma [code]: \Word.the_int 1 = 1\ by transfer simp lemma [code]: \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_add) lemma [code]: \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ for w :: \'a::len word\ by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) lemma [code]: \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_diff) lemma [code]: \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_mult) subsubsection \Basic conversions\ abbreviation word_of_nat :: \nat \ 'a::len word\ where \word_of_nat \ of_nat\ abbreviation word_of_int :: \int \ 'a::len word\ where \word_of_int \ of_int\ lemma word_of_nat_eq_iff: \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_eq_iff: \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_eq_0_iff: \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) lemma word_of_int_eq_0_iff: \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) context semiring_1 begin lift_definition unsigned :: \'b::len word \ 'a\ is \of_nat \ nat \ take_bit LENGTH('b)\ by simp lemma unsigned_0 [simp]: \unsigned 0 = 0\ by transfer simp lemma unsigned_1 [simp]: \unsigned 1 = 1\ by transfer simp lemma unsigned_numeral [simp]: \unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\ by transfer (simp add: nat_take_bit_eq) lemma unsigned_neg_numeral [simp]: \unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\ by transfer simp end context semiring_1 begin lemma unsigned_of_nat: \unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\ by transfer (simp add: nat_eq_iff take_bit_of_nat) lemma unsigned_of_int: \unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\ by transfer simp end context semiring_char_0 begin lemma unsigned_word_eqI: \v = w\ if \unsigned v = unsigned w\ using that by transfer (simp add: eq_nat_nat_iff) lemma word_eq_iff_unsigned: \v = w \ unsigned v = unsigned w\ by (auto intro: unsigned_word_eqI) lemma inj_unsigned [simp]: \inj unsigned\ by (rule injI) (simp add: unsigned_word_eqI) lemma unsigned_eq_0_iff: \unsigned w = 0 \ w = 0\ using word_eq_iff_unsigned [of w 0] by simp end context ring_1 begin lift_definition signed :: \'b::len word \ 'a\ is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ by (simp flip: signed_take_bit_decr_length_iff) lemma signed_0 [simp]: \signed 0 = 0\ by transfer simp lemma signed_1 [simp]: \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ by (transfer fixing: uminus; cases \LENGTH('b)\) (auto dest: gr0_implies_Suc) lemma signed_minus_1 [simp]: \signed (- 1 :: 'b::len word) = - 1\ by (transfer fixing: uminus) simp lemma signed_numeral [simp]: \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ by transfer simp lemma signed_neg_numeral [simp]: \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ by transfer simp lemma signed_of_nat: \signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\ by transfer simp lemma signed_of_int: \signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\ by transfer simp end context ring_char_0 begin lemma signed_word_eqI: \v = w\ if \signed v = signed w\ using that by transfer (simp flip: signed_take_bit_decr_length_iff) lemma word_eq_iff_signed: \v = w \ signed v = signed w\ by (auto intro: signed_word_eqI) lemma inj_signed [simp]: \inj signed\ by (rule injI) (simp add: signed_word_eqI) lemma signed_eq_0_iff: \signed w = 0 \ w = 0\ using word_eq_iff_signed [of w 0] by simp end abbreviation unat :: \'a::len word \ nat\ where \unat \ unsigned\ abbreviation uint :: \'a::len word \ int\ where \uint \ unsigned\ abbreviation sint :: \'a::len word \ int\ where \sint \ signed\ abbreviation ucast :: \'a::len word \ 'b::len word\ where \ucast \ unsigned\ abbreviation scast :: \'a::len word \ 'b::len word\ where \scast \ signed\ context includes lifting_syntax begin lemma [transfer_rule]: \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ using unsigned.transfer [where ?'a = nat] by simp lemma [transfer_rule]: \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ using unsigned.transfer [where ?'a = int] by (simp add: comp_def) lemma [transfer_rule]: \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ using signed.transfer [where ?'a = int] by simp lemma [transfer_rule]: \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ by simp qed lemma [transfer_rule]: \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ by simp qed end lemma of_nat_unat [simp]: \of_nat (unat w) = unsigned w\ by transfer simp lemma of_int_uint [simp]: \of_int (uint w) = unsigned w\ by transfer simp lemma of_int_sint [simp]: \of_int (sint a) = signed a\ by transfer (simp_all add: take_bit_signed_take_bit) lemma nat_uint_eq [simp]: \nat (uint w) = unat w\ by transfer simp lemma sgn_uint_eq [simp]: \sgn (uint w) = of_bool (w \ 0)\ by transfer (simp add: less_le) text \Aliasses only for code generation\ context begin qualified lift_definition of_int :: \int \ 'a::len word\ is \take_bit LENGTH('a)\ . qualified lift_definition of_nat :: \nat \ 'a::len word\ is \int \ take_bit LENGTH('a)\ . qualified lift_definition the_nat :: \'a::len word \ nat\ is \nat \ take_bit LENGTH('a)\ by simp qualified lift_definition the_signed_int :: \'a::len word \ int\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) qualified lift_definition cast :: \'a::len word \ 'b::len word\ is \take_bit LENGTH('a)\ by simp qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) end lemma [code_abbrev, simp]: \Word.the_int = uint\ by transfer rule lemma [code]: \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer simp lemma [code_abbrev, simp]: \Word.of_int = word_of_int\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ by transfer (simp add: take_bit_of_nat) lemma [code_abbrev, simp]: \Word.of_nat = word_of_nat\ by (rule; transfer) (simp add: take_bit_of_nat) lemma [code]: \Word.the_nat w = nat (Word.the_int w)\ by transfer simp lemma [code_abbrev, simp]: \Word.the_nat = unat\ by (rule; transfer) simp lemma [code]: \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: signed_take_bit_take_bit) lemma [code_abbrev, simp]: \Word.the_signed_int = sint\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.cast = ucast\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.signed_cast = scast\ by (rule; transfer) simp lemma [code]: \unsigned w = of_nat (nat (Word.the_int w))\ by transfer simp lemma [code]: \signed w = of_int (Word.the_signed_int w)\ by transfer simp subsubsection \Basic ordering\ instantiation word :: (len) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by (standard; transfer) auto end interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ by (standard; transfer) simp lemma word_of_nat_less_eq_iff: \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_eq_iff: \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_less_iff: \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_iff: \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule lemma word_le_def [code]: "a \ b \ uint a \ uint b" by transfer rule lemma word_less_def [code]: "a < b \ uint a < uint b" by transfer rule lemma word_greater_zero_iff: \a > 0 \ a \ 0\ for a :: \'a::len word\ by transfer (simp add: less_le) lemma of_nat_word_less_eq_iff: \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_nat_word_less_iff: \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_int_word_less_eq_iff: \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma of_int_word_less_iff: \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule subsection \Enumeration\ lemma inj_on_word_of_nat: \inj_on (word_of_nat :: nat \ 'a::len word) {0..<2 ^ LENGTH('a)}\ by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self) lemma UNIV_word_eq_word_of_nat: \(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\ (is \_ = ?A\) proof show \word_of_nat ` {0..<2 ^ LENGTH('a)} \ UNIV\ by simp show \UNIV \ ?A\ proof fix w :: \'a word\ show \w \ (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\ by (rule image_eqI [of _ _ \unat w\]; transfer) simp_all qed qed instantiation word :: (len) enum begin definition enum_word :: \'a word list\ where \enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\ definition enum_all_word :: \('a word \ bool) \ bool\ where \enum_all_word = Ball UNIV\ definition enum_ex_word :: \('a word \ bool) \ bool\ where \enum_ex_word = Bex UNIV\ lemma [code]: \Enum.enum_all P \ Ball UNIV P\ \Enum.enum_ex P \ Bex UNIV P\ for P :: \'a word \ bool\ by (simp_all add: enum_all_word_def enum_ex_word_def) instance by standard (simp_all add: UNIV_word_eq_word_of_nat inj_on_word_of_nat enum_word_def enum_all_word_def enum_ex_word_def distinct_map) end subsection \Bit-wise operations\ instantiation word :: (len) semiring_modulo begin lift_definition divide_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ by simp lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ by simp instance proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed end instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) qed lemma word_bit_induct [case_names zero even odd]: \P a\ if word_zero: \P 0\ and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ for P and a :: \'a::len word\ proof - define m :: nat where \m = LENGTH('a) - Suc 0\ then have l: \LENGTH('a) = Suc m\ by simp define n :: nat where \n = unat a\ then have \n < 2 ^ LENGTH('a)\ by transfer (simp add: take_bit_eq_mod) then have \n < 2 * 2 ^ m\ by (simp add: l) then have \P (of_nat n)\ proof (induction n rule: nat_bit_induct) case zero show ?case by simp (rule word_zero) next case (even n) then have \n < 2 ^ m\ by simp with even.IH have \P (of_nat n)\ by simp moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ by (auto simp add: word_greater_zero_iff l word_of_nat_eq_0_iff) moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (2 * of_nat n)\ by (rule word_even) then show ?case by simp next case (odd n) then have \Suc n \ 2 ^ m\ by simp with odd.IH have \P (of_nat n)\ by simp moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (1 + 2 * of_nat n)\ by (rule word_odd) then show ?case by simp qed moreover have \of_nat (nat (uint a)) = a\ by transfer simp ultimately show ?thesis by (simp add: n_def) qed lemma bit_word_half_eq: \(of_bool b + a * 2) div 2 = a\ if \a < 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof (cases \2 \ LENGTH('a::len)\) case False have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int by auto with False that show ?thesis by transfer (simp add: eq_iff) next case True obtain n where length: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all show ?thesis proof (cases b) case False moreover have \a * 2 div 2 = a\ using that proof transfer fix k :: int from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ by simp moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ - with \LENGTH('a) = Suc n\ - have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ - by (simp add: take_bit_eq_mod divmod_digit_0) + with \LENGTH('a) = Suc n\ have \take_bit LENGTH('a) k = take_bit n k\ + by (auto simp add: take_bit_Suc_from_most) ultimately have \take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by simp qed ultimately show ?thesis by simp next case True moreover have \(1 + a * 2) div 2 = a\ using that proof transfer fix k :: int from length have \(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\ using pos_zmod_mult_2 [of \2 ^ n\ k] by (simp add: ac_simps) moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ - with \LENGTH('a) = Suc n\ - have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ - by (simp add: take_bit_eq_mod divmod_digit_0) + with \LENGTH('a) = Suc n\ have \take_bit LENGTH('a) k = take_bit n k\ + by (auto simp add: take_bit_Suc_from_most) ultimately have \take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by (auto simp add: take_bit_Suc) qed ultimately show ?thesis by simp qed qed lemma even_mult_exp_div_word_iff: \even (a * 2 ^ m div 2 ^ n) \ \ ( m \ n \ n < LENGTH('a) \ odd (a div 2 ^ (n - m)))\ for a :: \'a::len word\ by transfer (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) instantiation word :: (len) semiring_bits begin lift_definition bit_word :: \'a word \ nat \ bool\ is \\k n. n < LENGTH('a) \ bit k n\ proof fix k l :: int and n :: nat assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ proof (cases \n < LENGTH('a)\) case True from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ by simp then show ?thesis by (simp add: bit_take_bit_iff) next case False then show ?thesis by simp qed qed instance proof show \P a\ if stable: \\a. a div 2 = a \ P a\ and rec: \\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)\ for P and a :: \'a word\ proof (induction a rule: word_bit_induct) case zero have \0 div 2 = (0::'a word)\ by transfer simp with stable [of 0] show ?case by simp next case (even a) with rec [of a False] show ?case using bit_word_half_eq [of a False] by (simp add: ac_simps) next case (odd a) with rec [of a True] show ?case using bit_word_half_eq [of a True] by (simp add: ac_simps) qed show \bit a n \ odd (a div 2 ^ n)\ for a :: \'a word\ and n by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) show \0 div a = 0\ for a :: \'a word\ by transfer simp show \a div 1 = a\ for a :: \'a word\ by transfer simp have \
: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" by (metis le_less take_bit_eq_mod take_bit_nonnegative) have less_power: "\n i p. (i::int) mod numeral p ^ n < numeral p ^ n" by simp show \a mod b div b = 0\ for a b :: \'a word\ apply transfer apply (simp add: take_bit_eq_mod mod_eq_0_iff_dvd dvd_def) by (metis (no_types, opaque_lifting) "\
" Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign le_less_trans mult_eq_0_iff take_bit_eq_mod take_bit_nonnegative zdiv_eq_0_iff zmod_le_nonneg_dividend) show \(1 + a) div 2 = a div 2\ if \even a\ for a :: \'a word\ using that by transfer (auto dest: le_Suc_ex simp add: take_bit_Suc elim!: evenE) show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat by transfer (simp, simp add: exp_div_exp_eq) show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" for a :: "'a word" and m n :: nat apply transfer apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) apply (simp add: drop_bit_take_bit) done show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" for a :: "'a word" and m n :: nat by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ if \m \ n\ for a :: "'a word" and m n :: nat using that apply transfer apply (auto simp flip: take_bit_eq_mod) apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) done show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ for a :: "'a word" and m n :: nat by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ for m n :: nat by transfer (simp flip: drop_bit_eq_div mask_eq_exp_minus_1 add: bit_simps even_drop_bit_iff_not_bit not_less) show \even (a * 2 ^ m div 2 ^ n) \ n < m \ (2::'a word) ^ n = 0 \ m \ n \ even (a div 2 ^ (n - m))\ for a :: \'a word\ and m n :: nat proof transfer show \even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \ n < m \ take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 \ (m \ n \ even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\ for m n :: nat and k l :: int by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) qed qed end lemma bit_word_eqI: \a = b\ if \\n. n < LENGTH('a) \ bit a n \ bit b n\ for a b :: \'a::len word\ using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) lemma bit_imp_le_length: \n < LENGTH('a)\ if \bit w n\ for w :: \'a::len word\ using that by transfer simp lemma not_bit_length [simp]: \\ bit w LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma finite_bit_word [simp]: \finite {n. bit w n}\ for w :: \'a::len word\ proof - have \{n. bit w n} \ {0..LENGTH('a)}\ by (auto dest: bit_imp_le_length) moreover have \finite {0..LENGTH('a)}\ by simp ultimately show ?thesis by (rule finite_subset) qed lemma bit_numeral_word_iff [simp]: \bit (numeral w :: 'a::len word) n \ n < LENGTH('a) \ bit (numeral w :: int) n\ by transfer simp lemma bit_neg_numeral_word_iff [simp]: \bit (- numeral w :: 'a::len word) n \ n < LENGTH('a) \ bit (- numeral w :: int) n\ by transfer simp instantiation word :: (len) ring_bit_operations begin lift_definition not_word :: \'a word \ 'a word\ is not by (simp add: take_bit_not_iff) lift_definition and_word :: \'a word \ 'a word \ 'a word\ is \and\ by simp lift_definition or_word :: \'a word \ 'a word \ 'a word\ is or by simp lift_definition xor_word :: \'a word \ 'a word \ 'a word\ is xor by simp lift_definition mask_word :: \nat \ 'a word\ is mask . lift_definition set_bit_word :: \nat \ 'a word \ 'a word\ is set_bit by (simp add: set_bit_def) lift_definition unset_bit_word :: \nat \ 'a word \ 'a word\ is unset_bit by (simp add: unset_bit_def) lift_definition flip_bit_word :: \nat \ 'a word \ 'a word\ is flip_bit by (simp add: flip_bit_def) lift_definition push_bit_word :: \nat \ 'a word \ 'a word\ is push_bit proof - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ by simp moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lift_definition drop_bit_word :: \nat \ 'a word \ 'a word\ is \\n. drop_bit n \ take_bit LENGTH('a)\ by (simp add: take_bit_eq_mod) lift_definition take_bit_word :: \nat \ 'a word \ 'a word\ is \\n. take_bit (min LENGTH('a) n)\ by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) instance apply (standard; transfer) apply (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 bit_simps set_bit_def flip_bit_def take_bit_drop_bit simp flip: drop_bit_eq_div take_bit_eq_mod) apply (simp_all add: drop_bit_take_bit flip: push_bit_eq_mult) done end lemma [code]: \push_bit n w = w * 2 ^ n\ for w :: \'a::len word\ by (fact push_bit_eq_mult) lemma [code]: \Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\ by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) lemma [code]: \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: not_le not_less ac_simps min_absorb2) lemma [code_abbrev]: \push_bit n 1 = (2 :: 'a::len word) ^ n\ by (fact push_bit_of_1) context includes bit_operations_syntax begin lemma [code]: \NOT w = Word.of_int (NOT (Word.the_int w))\ for w :: \'a::len word\ by transfer (simp add: take_bit_not_take_bit) lemma [code]: \Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp lemma [code]: \set_bit n w = w OR push_bit n 1\ for w :: \'a::len word\ by (fact set_bit_eq_or) lemma [code]: \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: \'a::len word\ by (fact unset_bit_eq_and_not) lemma [code]: \flip_bit n w = w XOR push_bit n 1\ for w :: \'a::len word\ by (fact flip_bit_eq_xor) context includes lifting_syntax begin lemma set_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) set_bit set_bit\ by (unfold set_bit_def) transfer_prover lemma unset_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\ by (unfold unset_bit_def) transfer_prover lemma flip_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\ by (unfold flip_bit_def) transfer_prover lemma signed_take_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) (\n k. signed_take_bit n (take_bit LENGTH('a::len) k)) (signed_take_bit :: nat \ 'a word \ 'a word)\ proof - let ?K = \\n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \ bit k n) * NOT (mask n)\ let ?W = \\n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\ have \((=) ===> pcr_word ===> pcr_word) ?K ?W\ by transfer_prover also have \?K = (\n k. signed_take_bit n (take_bit LENGTH('a::len) k))\ by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps) also have \?W = signed_take_bit\ by (simp add: fun_eq_iff signed_take_bit_def) finally show ?thesis . qed end end subsection \Conversions including casts\ subsubsection \Generic unsigned conversion\ context semiring_bits begin lemma bit_unsigned_iff [bit_simps]: \bit (unsigned w) n \ possible_bit TYPE('a) n \ bit w n\ for w :: \'b::len word\ by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff) end lemma possible_bit_word[simp]: \possible_bit TYPE(('a :: len) word) m \ m < LENGTH('a)\ by (simp add: possible_bit_def linorder_not_le) context semiring_bit_operations begin lemma unsigned_minus_1_eq_mask: \unsigned (- 1 :: 'b::len word) = mask LENGTH('b)\ by (transfer fixing: mask) (simp add: nat_mask_eq of_nat_mask_eq) lemma unsigned_push_bit_eq: \unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\ for w :: \'b::len word\ proof (rule bit_eqI) fix m assume \possible_bit TYPE('a) m\ show \bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\ proof (cases \n \ m\) case True with \possible_bit TYPE('a) m\ have \possible_bit TYPE('a) (m - n)\ by (simp add: possible_bit_less_imp) with True show ?thesis by (simp add: bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff not_le ac_simps) next case False then show ?thesis by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff) qed qed lemma unsigned_take_bit_eq: \unsigned (take_bit n w) = take_bit n (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Bit_Operations.bit_take_bit_iff) end context unique_euclidean_semiring_with_bit_operations begin lemma unsigned_drop_bit_eq: \unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\ for w :: \'b::len word\ by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Bit_Operations.bit_drop_bit_eq possible_bit_def dest: bit_imp_le_length) end lemma ucast_drop_bit_eq: \ucast (drop_bit n w) = drop_bit n (ucast w :: 'b::len word)\ if \LENGTH('a) \ LENGTH('b)\ for w :: \'a::len word\ by (rule bit_word_eqI) (use that in \auto simp add: bit_unsigned_iff bit_drop_bit_eq dest: bit_imp_le_length\) context semiring_bit_operations begin context includes bit_operations_syntax begin lemma unsigned_and_eq: \unsigned (v AND w) = unsigned v AND unsigned w\ for v w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps) lemma unsigned_or_eq: \unsigned (v OR w) = unsigned v OR unsigned w\ for v w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps) lemma unsigned_xor_eq: \unsigned (v XOR w) = unsigned v XOR unsigned w\ for v w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps) end end context ring_bit_operations begin context includes bit_operations_syntax begin lemma unsigned_not_eq: \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ for w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps) end end context unique_euclidean_semiring_numeral begin lemma unsigned_greater_eq [simp]: \0 \ unsigned w\ for w :: \'b::len word\ by (transfer fixing: less_eq) simp lemma unsigned_less [simp]: \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ by (transfer fixing: less) simp end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end subsubsection \Generic signed conversion\ context ring_bit_operations begin lemma bit_signed_iff [bit_simps]: \bit (signed w) n \ possible_bit TYPE('a) n \ bit w (min (LENGTH('b) - Suc 0) n)\ for w :: \'b::len word\ by (transfer fixing: bit) (auto simp add: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def) lemma signed_push_bit_eq: \signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\ for w :: \'b::len word\ apply (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj) apply (cases n, simp_all add: min_def) done lemma signed_take_bit_eq: \signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\ for w :: \'b::len word\ by (transfer fixing: take_bit; cases \LENGTH('b)\) (auto simp add: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq) context includes bit_operations_syntax begin lemma signed_not_eq: \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ for w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj) (auto simp: min_def) lemma signed_and_eq: \signed (v AND w) = signed v AND signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff) lemma signed_or_eq: \signed (v OR w) = signed v OR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff) lemma signed_xor_eq: \signed (v XOR w) = signed v XOR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff) end end subsubsection \More\ lemma sint_greater_eq: \- (2 ^ (LENGTH('a) - Suc 0)) \ sint w\ for w :: \'a::len word\ proof (cases \bit w (LENGTH('a) - Suc 0)\) case True then show ?thesis by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps) next have *: \- (2 ^ (LENGTH('a) - Suc 0)) \ (0::int)\ by simp case False then show ?thesis by transfer (auto simp add: signed_take_bit_eq intro: order_trans *) qed lemma sint_less: \sint w < 2 ^ (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ by (cases \bit w (LENGTH('a) - Suc 0)\; transfer) (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper) lemma unat_div_distrib: \unat (v div w) = unat v div unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule div_le_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k div (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff) qed lemma unat_mod_distrib: \unat (v mod w) = unat v mod unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule mod_less_eq_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k mod (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff) qed lemma uint_div_distrib: \uint (v div w) = uint v div uint w\ proof - have \int (unat (v div w)) = int (unat v div unat w)\ by (simp add: unat_div_distrib) then show ?thesis by (simp add: of_nat_div) qed lemma unat_drop_bit_eq: \unat (drop_bit n w) = drop_bit n (unat w)\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq) lemma uint_mod_distrib: \uint (v mod w) = uint v mod uint w\ proof - have \int (unat (v mod w)) = int (unat v mod unat w)\ by (simp add: unat_mod_distrib) then show ?thesis by (simp add: of_nat_mod) qed context semiring_bit_operations begin lemma unsigned_ucast_eq: \unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff not_le) end context ring_bit_operations begin lemma signed_ucast_eq: \signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\ for w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps min_less_iff_disj) lemma signed_scast_eq: \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ for w :: \'b::len word\ by (simp add: bit_eq_iff bit_simps min_less_iff_disj) end lemma uint_nonnegative: "0 \ uint w" by (fact unsigned_greater_eq) lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len word" by (fact unsigned_less) lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len word" by transfer (simp add: take_bit_eq_mod) lemma word_uint_eqI: "uint a = uint b \ a = b" by (fact unsigned_word_eqI) lemma word_uint_eq_iff: "a = b \ uint a = uint b" by (fact word_eq_iff_unsigned) lemma uint_word_of_int_eq: \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer rule lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" by (simp add: uint_word_of_int_eq take_bit_eq_mod) lemma word_of_int_uint: "word_of_int (uint w) = w" by transfer simp lemma word_div_def [code]: "a div b = word_of_int (uint a div uint b)" by transfer rule lemma word_mod_def [code]: "a mod b = word_of_int (uint a mod uint b)" by transfer rule lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" proof fix x :: "'a word" assume "\x. PROP P (word_of_int x)" then have "PROP P (word_of_int (uint x))" . then show "PROP P x" by (simp only: word_of_int_uint) qed lemma sint_uint: \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ for w :: \'a::len word\ by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) lemma unat_eq_nat_uint: \unat w = nat (uint w)\ by simp lemma ucast_eq: \ucast w = word_of_int (uint w)\ by transfer simp lemma scast_eq: \scast w = word_of_int (sint w)\ by transfer simp lemma uint_0_eq: \uint 0 = 0\ by (fact unsigned_0) lemma uint_1_eq: \uint 1 = 1\ by (fact unsigned_1) lemma word_m1_wi: "- 1 = word_of_int (- 1)" by simp lemma uint_0_iff: "uint x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0_iff: "unat x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0: "unat 0 = 0" by (fact unsigned_0) lemma unat_gt_0: "0 < unat x \ x \ 0" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0: "ucast 0 = 0" by (fact unsigned_0) lemma sint_0: "sint 0 = 0" by (fact signed_0) lemma scast_0: "scast 0 = 0" by (fact signed_0) lemma sint_n1: "sint (- 1) = - 1" by (fact signed_minus_1) lemma scast_n1: "scast (- 1) = - 1" by (fact signed_minus_1) lemma uint_1: "uint (1::'a::len word) = 1" by (fact uint_1_eq) lemma unat_1: "unat (1::'a::len word) = 1" by (fact unsigned_1) lemma ucast_1: "ucast (1::'a::len word) = 1" by (fact unsigned_1) instantiation word :: (len) size begin lift_definition size_word :: \'a word \ nat\ is \\_. LENGTH('a)\ .. instance .. end lemma word_size [code]: \size w = LENGTH('a)\ for w :: \'a::len word\ by (fact size_word.rep_eq) lemma word_size_gt_0 [iff]: "0 < size w" for w :: "'a::len word" by (simp add: word_size) lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemma lens_not_0 [iff]: \size w \ 0\ for w :: \'a::len word\ by auto lift_definition source_size :: \('a::len word \ 'b) \ nat\ is \\_. LENGTH('a)\ . lift_definition target_size :: \('a \ 'b::len word) \ nat\ is \\_. LENGTH('b)\ .. lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lemma is_up_eq: \is_up f \ source_size f \ target_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) lemma is_down_eq: \is_down f \ target_size f \ source_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ is \\f. f \ take_bit LENGTH('a)\ by simp lemma word_int_case_eq_uint [code]: \word_int_case f w = f (uint w)\ by transfer simp translations "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" subsection \Arithmetic operations\ lemma div_word_self: \w div w = 1\ if \w \ 0\ for w :: \'a::len word\ using that by transfer simp lemma mod_word_self [simp]: \w mod w = 0\ for w :: \'a::len word\ apply (cases \w = 0\) apply auto using div_mult_mod_eq [of w w] by (simp add: div_word_self) lemma div_word_less: \w div v = 0\ if \w < v\ for w v :: \'a::len word\ using that by transfer simp lemma mod_word_less: \w mod v = w\ if \w < v\ for w v :: \'a::len word\ using div_mult_mod_eq [of w v] using that by (simp add: div_word_less) lemma div_word_one [simp]: \1 div w = of_bool (w = 1)\ for w :: \'a::len word\ proof transfer fix k :: int show \take_bit LENGTH('a) (take_bit LENGTH('a) 1 div take_bit LENGTH('a) k) = take_bit LENGTH('a) (of_bool (take_bit LENGTH('a) k = take_bit LENGTH('a) 1))\ proof (cases \take_bit LENGTH('a) k > 1\) case False with take_bit_nonnegative [of \LENGTH('a)\ k] have \take_bit LENGTH('a) k = 0 \ take_bit LENGTH('a) k = 1\ by linarith then show ?thesis by auto next case True then show ?thesis by simp qed qed lemma mod_word_one [simp]: \1 mod w = 1 - w * of_bool (w = 1)\ for w :: \'a::len word\ using div_mult_mod_eq [of 1 w] by auto lemma div_word_by_minus_1_eq [simp]: \w div - 1 = of_bool (w = - 1)\ for w :: \'a::len word\ by (auto intro: div_word_less simp add: div_word_self word_order.not_eq_extremum) lemma mod_word_by_minus_1_eq [simp]: \w mod - 1 = w * of_bool (w < - 1)\ for w :: \'a::len word\ proof (cases \w = - 1\) case True then show ?thesis by simp next case False moreover have \w < - 1\ using False by (simp add: word_order.not_eq_extremum) ultimately show ?thesis by (simp add: mod_word_less) qed text \Legacy theorems:\ lemma word_add_def [code]: "a + b = word_of_int (uint a + uint b)" by transfer (simp add: take_bit_add) lemma word_sub_wi [code]: "a - b = word_of_int (uint a - uint b)" by transfer (simp add: take_bit_diff) lemma word_mult_def [code]: "a * b = word_of_int (uint a * uint b)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_minus_def [code]: "- a = word_of_int (- uint a)" by transfer (simp add: take_bit_minus) lemma word_0_wi: "0 = word_of_int 0" by transfer simp lemma word_1_wi: "1 = word_of_int 1" by transfer simp lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lemma word_succ_alt [code]: "word_succ a = word_of_int (uint a + 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_pred_alt [code]: "word_pred a = word_of_int (uint a - 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemmas word_arith_wis = word_add_def word_sub_wi word_mult_def word_minus_def word_succ_alt word_pred_alt word_0_wi word_1_wi lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" by (transfer, simp)+ lemmas wi_hom_syms = wi_homs [symmetric] lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] lemma double_eq_zero_iff: \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof - define n where \n = LENGTH('a) - Suc 0\ then have *: \LENGTH('a) = Suc n\ by simp have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ using that by transfer (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ by transfer simp then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ by (simp add: *) ultimately show ?thesis by auto qed subsection \Ordering\ lift_definition word_sle :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k \ signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition word_sless :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) notation word_sle ("'(\s')") and word_sle ("(_/ \s _)" [51, 51] 50) and word_sless ("'(a <=s b \ sint a \ sint b\ by transfer simp lemma [code]: \a sint a < sint b\ by transfer simp lemma signed_ordering: \ordering word_sle word_sless\ apply (standard; transfer) using signed_take_bit_decr_length_iff by force+ lemma signed_linorder: \class.linorder word_sle word_sless\ by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) interpretation signed: linorder word_sle word_sless by (fact signed_linorder) lemma word_sless_eq: \x x <=s y \ x \ y\ by (fact signed.less_le) lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len word" by (fact word_coorder.extremum) lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by transfer (simp add: mask_eq_exp_minus_1) lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len word" by (fact word_order.extremum) lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y \ 0 \ y" for y :: "'a::len word" by (simp add: less_le) lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "a sint a < sint b" by transfer simp lemma word_le_nat_alt: "a \ b \ unat a \ unat b" by transfer (simp add: nat_le_eq_zle) lemma word_less_nat_alt: "a < b \ unat a < unat b" by transfer (auto simp add: less_le [of 0]) lemmas unat_mono = word_less_nat_alt [THEN iffD1] instance word :: (len) wellorder proof fix P :: "'a word \ bool" and a assume *: "(\b. (\a. a < b \ P a) \ P b)" have "wf (measure unat)" .. moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" by (auto simp add: word_less_nat_alt) ultimately have "wf {(a, b :: ('a::len) word). a < b}" by (rule wf_subset) then show "P a" using * by induction blast qed lemma wi_less: "(word_of_int n < (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) lemma wi_le: "(word_of_int n \ (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) subsection \Bit-wise operations\ context includes bit_operations_syntax begin lemma uint_take_bit_eq: \uint (take_bit n w) = take_bit n (uint w)\ by transfer (simp add: ac_simps) lemma take_bit_word_eq_self: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp lemma take_bit_length_eq [simp]: \take_bit LENGTH('a) w = w\ for w :: \'a::len word\ by (rule take_bit_word_eq_self) simp lemma bit_word_of_int_iff: \bit (word_of_int k :: 'a::len word) n \ n < LENGTH('a) \ bit k n\ by transfer rule lemma bit_uint_iff: \bit (uint w) n \ n < LENGTH('a) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff) lemma bit_sint_iff: \bit (sint w) n \ n \ LENGTH('a) \ bit w (LENGTH('a) - 1) \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff min_def le_less not_less) lemma bit_word_ucast_iff: \bit (ucast w :: 'b::len word) n \ n < LENGTH('a) \ n < LENGTH('b) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff ac_simps) lemma bit_word_scast_iff: \bit (scast w :: 'b::len word) n \ n < LENGTH('b) \ (bit w n \ LENGTH('a) \ n \ bit w (LENGTH('a) - Suc 0))\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff le_less min_def) lemma bit_word_iff_drop_bit_and [code]: \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) lemma word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" by (transfer, simp add: take_bit_not_take_bit)+ definition even_word :: \'a::len word \ bool\ where [code_abbrev]: \even_word = even\ lemma even_word_iff [code]: \even_word a \ a AND 1 = 0\ by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def) lemma map_bit_range_eq_if_take_bit_eq: \map (bit k) [0.. if \take_bit n k = take_bit n l\ for k l :: int using that proof (induction n arbitrary: k l) case 0 then show ?case by simp next case (Suc n) from Suc.prems have \take_bit n (k div 2) = take_bit n (l div 2)\ by (simp add: take_bit_Suc) then have \map (bit (k div 2)) [0.. by (rule Suc.IH) moreover have \bit (r div 2) = bit r \ Suc\ for r :: int by (simp add: fun_eq_iff bit_Suc) moreover from Suc.prems have \even k \ even l\ by (auto simp add: take_bit_Suc elim!: evenE oddE) arith+ ultimately show ?case by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) (simp add: bit_0) qed lemma take_bit_word_Bit0_eq [simp]: \take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (numeral m)\ (is ?P) and take_bit_word_Bit1_eq [simp]: \take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (numeral m)\ (is ?Q) and take_bit_word_minus_Bit0_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (- numeral m)\ (is ?R) and take_bit_word_minus_Bit1_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\ (is ?S) proof - define w :: \'a::len word\ where \w = numeral m\ moreover define q :: nat where \q = pred_numeral n\ ultimately have num: \numeral m = w\ \numeral (num.Bit0 m) = 2 * w\ \numeral (num.Bit1 m) = 1 + 2 * w\ \numeral (Num.inc m) = 1 + w\ \pred_numeral n = q\ \numeral n = Suc q\ by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps numeral_inc numeral_eq_Suc flip: mult_2) have even: \take_bit (Suc q) (2 * w) = 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_take_bit_iff bit_double_iff) have odd: \take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_double_iff even_bit_succ_iff) show ?P using even [of w] by (simp add: num) show ?Q using odd [of w] by (simp add: num) show ?R using even [of \- w\] by (simp add: num) show ?S using odd [of \- (1 + w)\] by (simp add: num) qed subsection \More shift operations\ lift_definition signed_drop_bit :: \nat \ 'a word \ 'a::len word\ is \\n. drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)\ using signed_take_bit_decr_length_iff by (simp add: take_bit_drop_bit) force lemma bit_signed_drop_bit_iff [bit_simps]: \bit (signed_drop_bit m w) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else m + n)\ for w :: \'a::len word\ apply transfer apply (auto simp add: bit_drop_bit_eq bit_signed_take_bit_iff not_le min_def) apply (metis add.commute le_antisym less_diff_conv less_eq_decr_length_iff) apply (metis le_antisym less_eq_decr_length_iff) done lemma [code]: \Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ for w :: \'a::len word\ by transfer simp lemma signed_drop_bit_of_0 [simp]: \signed_drop_bit n 0 = 0\ by transfer simp lemma signed_drop_bit_of_minus_1 [simp]: \signed_drop_bit n (- 1) = - 1\ by transfer simp lemma signed_drop_bit_signed_drop_bit [simp]: \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ for w :: \'a::len word\ proof (cases \LENGTH('a)\) case 0 then show ?thesis using len_not_eq_0 by blast next case (Suc n) then show ?thesis by (force simp add: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps intro!: bit_word_eqI) qed lemma signed_drop_bit_0 [simp]: \signed_drop_bit 0 w = w\ by transfer (simp add: take_bit_signed_take_bit) lemma sint_signed_drop_bit_eq: \sint (signed_drop_bit n w) = drop_bit n (sint w)\ proof (cases \LENGTH('a) = 0 \ n=0\) case False then show ?thesis apply simp apply (rule bit_eqI) by (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) qed auto subsection \Single-bit operations\ lemma set_bit_eq_idem_iff: \Bit_Operations.set_bit n w = w \ bit w n \ n \ LENGTH('a)\ for w :: \'a::len word\ by (simp add: bit_eq_iff) (auto simp add: bit_simps not_le) lemma unset_bit_eq_idem_iff: \unset_bit n w = w \ bit w n \ n \ LENGTH('a)\ for w :: \'a::len word\ by (simp add: bit_eq_iff) (auto simp add: bit_simps dest: bit_imp_le_length) lemma flip_bit_eq_idem_iff: \flip_bit n w = w \ n \ LENGTH('a)\ for w :: \'a::len word\ using linorder_le_less_linear by (simp add: bit_eq_iff) (auto simp add: bit_simps) subsection \Rotation\ lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (n mod LENGTH('a)) k)\ subgoal for n k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) done lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (n mod LENGTH('a)) (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ subgoal for n k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) done lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) (take_bit (nat (r mod int LENGTH('a))) k)\ subgoal for r k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) done lemma word_rotl_eq_word_rotr [code]: \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all lemma word_roti_eq_word_rotr_word_rotl [code]: \word_roti i w = (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ proof (cases \i \ 0\) case True moreover define n where \n = nat i\ ultimately have \i = int n\ by simp moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: nat_mod_distrib) ultimately show ?thesis by simp next case False moreover define n where \n = nat (- i)\ ultimately have \i = - int n\ \n > 0\ by simp_all moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) ultimately show ?thesis by simp qed lemma bit_word_rotr_iff [bit_simps]: \bit (word_rotr m w) n \ n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ for w :: \'a::len word\ proof transfer fix k :: int and m n :: nat define q where \q = m mod LENGTH('a)\ have \q < LENGTH('a)\ by (simp add: q_def) then have \q \ LENGTH('a)\ by simp have \m mod LENGTH('a) = q\ by (simp add: q_def) moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ using \q < LENGTH('a)\ by (cases \q + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff le_mod_geq ac_simps) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (m mod LENGTH('a)) k)) n \ n < LENGTH('a) \ (n + m) mod LENGTH('a) < LENGTH('a) \ bit k ((n + m) mod LENGTH('a))\ by simp qed lemma bit_word_rotl_iff [bit_simps]: \bit (word_rotl m w) n \ n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ for w :: \'a::len word\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) lemma bit_word_roti_iff [bit_simps]: \bit (word_roti k w) n \ n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ for w :: \'a::len word\ proof transfer fix k l :: int and n :: nat define m where \m = nat (k mod int LENGTH('a))\ have \m < LENGTH('a)\ by (simp add: nat_less_iff m_def) then have \m \ LENGTH('a)\ by simp have \k mod int LENGTH('a) = int m\ by (simp add: nat_less_iff m_def) moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ using \m < LENGTH('a)\ by (cases \m + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff nat_less_iff not_le not_less ac_simps le_diff_conv le_mod_geq) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) (take_bit (nat (k mod int LENGTH('a))) l)) n \ n < LENGTH('a) \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) \ bit l (nat ((int n + k) mod int LENGTH('a)))\ by simp qed lemma uint_word_rotr_eq: \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (uint w)) (uint (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ by transfer (simp add: take_bit_concat_bit_eq) lemma [code]: \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (Word.the_int w)) (Word.the_int (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ using uint_word_rotr_eq [of n w] by simp subsection \Split and cat operations\ lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) lemma word_cat_eq: \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: concat_bit_eq ac_simps) lemma word_cat_eq' [code]: \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ for a :: \'a::len word\ and b :: \'b::len word\ by transfer (simp add: concat_bit_take_bit_eq) lemma bit_word_cat_iff [bit_simps]: \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) definition word_split :: \'a::len word \ 'b::len word \ 'c::len word\ where \word_split w = (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\ definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ subsection \More on conversions\ lemma int_word_sint: \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift) lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by (simp add: signed_of_int) lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)" for w :: "'a::len word" by transfer (simp add: take_bit_signed_take_bit) lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" by transfer (simp add: min_def) lemma wi_bintr: "LENGTH('a::len) \ n \ word_of_int (take_bit n w) = (word_of_int w :: 'a word)" by transfer simp lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" by (induct b, simp_all only: numeral.simps word_of_int_homs) declare word_numeral_alt [symmetric, code_abbrev] lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" by (simp only: word_numeral_alt wi_hom_neg) declare word_neg_numeral_alt [symmetric, code_abbrev] lemma uint_bintrunc [simp]: "uint (numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (numeral bin)" by transfer rule lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" by transfer rule lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" by transfer simp lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" by transfer simp lemma unat_bintrunc [simp]: "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" by transfer simp lemma unat_bintrunc_neg [simp]: "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" by transfer simp lemma size_0_eq: "size w = 0 \ v = w" for v w :: "'a::len word" by transfer simp lemma uint_ge_0 [iff]: "0 \ uint x" by (fact unsigned_greater_eq) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len word" by (fact unsigned_less) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using sint_greater_eq [of x] by simp lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using sint_less [of x] by simp lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len word" by (simp only: diff_less_0_iff_less uint_lt2p) lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" for x :: "'a::len word" by (simp only: not_le uint_m2p_neg) lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len word" using uint_bounded [of w] by (rule less_le_trans) simp lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) lemma uint_nat: "uint w = int (unat w)" by transfer simp lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma sint_numeral: "sint (numeral b :: 'a::len word) = (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" by (metis int_word_sint word_numeral_alt) lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" by (fact of_int_0) lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" by (fact of_int_1) lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" by (simp add: wi_hom_syms) lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" by (fact of_int_numeral) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" by (fact of_int_neg_numeral) lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" by transfer (simp add: take_bit_eq_mod) lemma word_int_split: "P (word_int_case f x) = (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" by transfer (auto simp add: take_bit_eq_mod) lemma word_int_split_asm: "P (word_int_case f x) = (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" by transfer (auto simp add: take_bit_eq_mod) lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" by transfer simp lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" by (simp add: word_size sint_greater_eq sint_less) lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" unfolding word_size by (rule less_le_trans [OF sint_lt]) lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" for w :: "'a::len word" unfolding word_size by (rule order_trans [OF _ sint_ge]) lemma word_unat_eq_iff: \v = w \ unat v = unat w\ for v w :: \'a::len word\ by (fact word_eq_iff_unsigned) subsection \Testing bits\ lemma bin_nth_uint_imp: "bit (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" by transfer (simp add: bit_take_bit_iff) lemma bin_nth_sint: "LENGTH('a) \ n \ bit (sint w) n = bit (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) lemma num_of_bintr': "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \take_bit LENGTH('a) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by simp qed lemma num_of_sbintr': "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by (simp add: take_bit_signed_take_bit) qed lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (take_bit (LENGTH('a::len)) (numeral x))" by transfer simp lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" by transfer (simp add: take_bit_signed_take_bit) text \ \cast\ -- note, no arg for new length, as it's determined by type of result, thus in \cast w = w\, the type means cast to length of \w\! \ lemma bit_ucast_iff: \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ bit a n\ by transfer (simp add: bit_take_bit_iff) lemma ucast_id [simp]: "ucast w = w" by transfer simp lemma scast_id [simp]: "scast w = w" by transfer (simp add: take_bit_signed_take_bit) lemma ucast_mask_eq: \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff) \ \literal u(s)cast\ lemma ucast_bintr [simp]: "ucast (numeral w :: 'a::len word) = word_of_int (take_bit (LENGTH('a)) (numeral w))" by transfer simp (* TODO: neg_numeral *) lemma scast_sbintr [simp]: "scast (numeral w ::'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))" by transfer simp lemma source_size: "source_size (c::'a::len word \ _) = LENGTH('a)" by transfer simp lemma target_size: "target_size (c::_ \ 'b::len word) = LENGTH('b)" by transfer simp lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up_down: \is_up c \ is_down d\ for c :: \'a::len word \ 'b::len word\ and d :: \'b::len word \ 'a::len word\ by transfer simp context fixes dummy_types :: \'a::len \ 'b::len\ begin private abbreviation (input) UCAST :: \'a::len word \ 'b::len word\ where \UCAST == ucast\ private abbreviation (input) SCAST :: \'a::len word \ 'b::len word\ where \SCAST == scast\ lemma down_cast_same: \UCAST = scast\ if \is_down UCAST\ by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit) lemma sint_up_scast: \sint (SCAST w) = sint w\ if \is_up SCAST\ using that by transfer (simp add: min_def Suc_leI le_diff_iff) lemma uint_up_ucast: \uint (UCAST w) = uint w\ if \is_up UCAST\ using that by transfer (simp add: min_def) lemma ucast_up_ucast: \ucast (UCAST w) = ucast w\ if \is_up UCAST\ using that by transfer (simp add: ac_simps) lemma ucast_up_ucast_id: \ucast (UCAST w) = w\ if \is_up UCAST\ using that by (simp add: ucast_up_ucast) lemma scast_up_scast: \scast (SCAST w) = scast w\ if \is_up SCAST\ using that by transfer (simp add: ac_simps) lemma scast_up_scast_id: \scast (SCAST w) = w\ if \is_up SCAST\ using that by (simp add: scast_up_scast) lemma isduu: \is_up UCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of UCAST d] by simp lemma isdus: \is_up SCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of SCAST d] by simp lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id] lemma up_ucast_surj: \surj (ucast :: 'b word \ 'a word)\ if \is_up UCAST\ by (rule surjI) (use that in \rule ucast_up_ucast_id\) lemma up_scast_surj: \surj (scast :: 'b word \ 'a word)\ if \is_up SCAST\ by (rule surjI) (use that in \rule scast_up_scast_id\) lemma down_ucast_inj: \inj_on UCAST A\ if \is_down (ucast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule ucast_down_ucast_id\) lemma down_scast_inj: \inj_on SCAST A\ if \is_down (scast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule scast_down_scast_id\) lemma ucast_down_wi: \UCAST (word_of_int x) = word_of_int x\ if \is_down UCAST\ using that by transfer simp lemma ucast_down_no: \UCAST (numeral bin) = numeral bin\ if \is_down UCAST\ using that by transfer simp end lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def lemma bit_last_iff: \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) for w :: \'a::len word\ proof - have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ by (simp add: bit_uint_iff) also have \\ \ ?Q\ by (simp add: sint_uint) finally show ?thesis . qed lemma drop_bit_eq_zero_iff_not_bit_last: \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ for w :: "'a::len word" proof (cases \LENGTH('a)\) case (Suc n) then show ?thesis apply transfer apply (simp add: take_bit_drop_bit) by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit odd_iff_mod_2_eq_one) qed auto lemma unat_div: \unat (x div y) = unat x div unat y\ by (fact unat_div_distrib) lemma unat_mod: \unat (x mod y) = unat x mod unat y\ by (fact unat_mod_distrib) subsection \Word Arithmetic\ lemmas less_eq_word_numeral_numeral [simp] = word_le_def [of \numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_numeral_numeral [simp] = word_less_def [of \numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_eq_word_minus_numeral_numeral [simp] = word_le_def [of \- numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_minus_numeral_numeral [simp] = word_less_def [of \- numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_eq_word_numeral_minus_numeral [simp] = word_le_def [of \numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_numeral_minus_numeral [simp] = word_less_def [of \numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_eq_word_minus_numeral_minus_numeral [simp] = word_le_def [of \- numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_minus_numeral_minus_numeral [simp] = word_less_def [of \- numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_numeral_minus_1 [simp] = word_less_def [of \numeral a\ \- 1\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas less_word_minus_numeral_minus_1 [simp] = word_less_def [of \- numeral a\ \- 1\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas sless_eq_word_numeral_numeral [simp] = word_sle_eq [of \numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_word_numeral_numeral [simp] = word_sless_alt [of \numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_eq_word_minus_numeral_numeral [simp] = word_sle_eq [of \- numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_word_minus_numeral_numeral [simp] = word_sless_alt [of \- numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_eq_word_numeral_minus_numeral [simp] = word_sle_eq [of \numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_word_numeral_minus_numeral [simp] = word_sless_alt [of \numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_eq_word_minus_numeral_minus_numeral [simp] = word_sle_eq [of \- numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas sless_word_minus_numeral_minus_numeral [simp] = word_sless_alt [of \- numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] for a b lemmas div_word_numeral_numeral [simp] = word_div_def [of \numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas div_word_minus_numeral_numeral [simp] = word_div_def [of \- numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas div_word_numeral_minus_numeral [simp] = word_div_def [of \numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas div_word_minus_numeral_minus_numeral [simp] = word_div_def [of \- numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas div_word_minus_1_numeral [simp] = word_div_def [of \- 1\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas div_word_minus_1_minus_numeral [simp] = word_div_def [of \- 1\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_numeral_numeral [simp] = word_mod_def [of \numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_minus_numeral_numeral [simp] = word_mod_def [of \- numeral a\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_numeral_minus_numeral [simp] = word_mod_def [of \numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_minus_numeral_minus_numeral [simp] = word_mod_def [of \- numeral a\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_minus_1_numeral [simp] = word_mod_def [of \- 1\ \numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemmas mod_word_minus_1_minus_numeral [simp] = word_mod_def [of \- 1\ \- numeral b\, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1] for a b lemma signed_drop_bit_of_1 [simp]: \signed_drop_bit n (1 :: 'a::len word) = of_bool (LENGTH('a) = 1 \ n = 0)\ apply (transfer fixing: n) apply (cases \LENGTH('a)\) apply (auto simp add: take_bit_signed_take_bit) apply (auto simp add: take_bit_drop_bit gr0_conv_Suc simp flip: take_bit_eq_self_iff_drop_bit_eq_0) done lemma take_bit_word_beyond_length_eq: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b lemma size_0_same': "size w = 0 \ w = v" for v w :: "'a::len word" by (unfold word_size) simp lemmas size_0_same = size_0_same' [unfolded word_size] lemmas unat_eq_0 = unat_0_iff lemmas unat_eq_zero = unat_0_iff lemma mask_1: "mask 1 = 1" by simp lemma mask_Suc_0: "mask (Suc 0) = 1" by simp lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd n" by simp lemma push_bit_word_beyond [simp]: \push_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by (transfer fixing: n) (simp add: take_bit_push_bit) lemma drop_bit_word_beyond [simp]: \drop_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by (transfer fixing: n) (simp add: drop_bit_take_bit) lemma signed_drop_bit_beyond: \signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\ if \LENGTH('a) \ n\ for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that) lemma take_bit_numeral_minus_numeral_word [simp]: \take_bit (numeral m) (- numeral n :: 'a::len word) = (case take_bit_num (numeral m) n of None \ 0 | Some q \ take_bit (numeral m) (2 ^ numeral m - numeral q))\ (is \?lhs = ?rhs\) proof (cases \LENGTH('a) \ numeral m\) case True then have *: \(take_bit (numeral m) :: 'a word \ 'a word) = id\ by (simp add: fun_eq_iff take_bit_word_eq_self) have **: \2 ^ numeral m = (0 :: 'a word)\ using True by (simp flip: exp_eq_zero_iff) show ?thesis by (auto simp only: * ** split: option.split dest!: take_bit_num_eq_None_imp [where ?'a = \'a word\] take_bit_num_eq_Some_imp [where ?'a = \'a word\]) simp_all next case False then show ?thesis by (transfer fixing: m n) simp qed lemma of_nat_inverse: \word_of_nat r = a \ r < 2 ^ LENGTH('a) \ unat a = r\ for a :: \'a::len word\ by (metis id_apply of_nat_eq_id take_bit_nat_eq_self_iff unsigned_of_nat) subsection \Transferring goals from words to ints\ lemma word_ths: shows word_succ_p1: "word_succ a = a + 1" and word_pred_m1: "word_pred a = a - 1" and word_pred_succ: "word_pred (word_succ a) = a" and word_succ_pred: "word_succ (word_pred a) = a" and word_mult_succ: "word_succ a * b = b + a * b" by (transfer, simp add: algebra_simps)+ lemma uint_cong: "x = y \ uint x = uint y" by simp lemma uint_word_ariths: fixes a b :: "'a::len word" shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) lemma uint_word_arith_bintrs: fixes a b :: "'a::len word" shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" by (simp_all add: uint_word_ariths take_bit_eq_mod) lemma sint_word_ariths: fixes a b :: "'a::len word" shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" subgoal by transfer (simp add: signed_take_bit_add) subgoal by transfer (simp add: signed_take_bit_diff) subgoal by transfer (simp add: signed_take_bit_mult) subgoal by transfer (simp add: signed_take_bit_minus) apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ) apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred) apply (simp_all add: sint_uint) done lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" unfolding word_pred_m1 by simp lemma succ_pred_no [simp]: "word_succ (numeral w) = numeral w + 1" "word_pred (numeral w) = numeral w - 1" "word_succ (- numeral w) = - numeral w + 1" "word_pred (- numeral w) = - numeral w - 1" by (simp_all add: word_succ_p1 word_pred_m1) lemma word_sp_01 [simp]: "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" by (simp_all add: word_succ_p1 word_pred_m1) \ \alternative approach to lifting arithmetic equalities\ lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp subsection \Order on fixed-length words\ lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp lemma udvd_iff_dvd: \x udvd y \ unat x dvd unat y\ by transfer (simp add: nat_dvd_iff) lemma udvd_iff_dvd_int: \v udvd w \ uint v dvd uint w\ by transfer rule lemma udvdI [intro]: \v udvd w\ if \unat w = unat v * unat u\ proof - from that have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvdE [elim]: fixes v w :: \'a::len word\ assumes \v udvd w\ obtains u :: \'a word\ where \unat w = unat v * unat u\ proof (cases \v = 0\) case True moreover from True \v udvd w\ have \w = 0\ by transfer simp ultimately show thesis using that by simp next case False then have \unat v > 0\ by (simp add: unat_gt_0) from \v udvd w\ have \unat v dvd unat w\ by (simp add: udvd_iff_dvd) then obtain n where \unat w = unat v * n\ .. moreover have \n < 2 ^ LENGTH('a)\ proof (rule ccontr) assume \\ n < 2 ^ LENGTH('a)\ then have \n \ 2 ^ LENGTH('a)\ by (simp add: not_le) then have \unat v * n \ 2 ^ LENGTH('a)\ using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] by simp with \unat w = unat v * n\ have \unat w \ 2 ^ LENGTH('a)\ by simp with unsigned_less [of w, where ?'a = nat] show False by linarith qed ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ by (auto simp add: take_bit_nat_eq_self_iff unsigned_of_nat intro: sym) with that show thesis . qed lemma udvd_imp_mod_eq_0: \w mod v = 0\ if \v udvd w\ using that by transfer simp lemma mod_eq_0_imp_udvd [intro?]: \v udvd w\ if \w mod v = 0\ proof - from that have \unat (w mod v) = unat 0\ by simp then have \unat w mod unat v = 0\ by (simp add: unat_mod_distrib) then have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvd_imp_dvd: \v dvd w\ if \v udvd w\ for v w :: \'a::len word\ proof - from that obtain u :: \'a word\ where \unat w = unat v * unat u\ .. then have \(word_of_nat (unat w) :: 'a word) = word_of_nat (unat v * unat u)\ by simp then have \w = v * u\ by simp then show \v dvd w\ .. qed lemma exp_dvd_iff_exp_udvd: \2 ^ n dvd w \ 2 ^ n udvd w\ for v w :: \'a::len word\ proof assume \2 ^ n udvd w\ then show \2 ^ n dvd w\ by (rule udvd_imp_dvd) next assume \2 ^ n dvd w\ then obtain u :: \'a word\ where \w = 2 ^ n * u\ .. then have \w = push_bit n u\ by (simp add: push_bit_eq_mult) then show \2 ^ n udvd w\ by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod) qed lemma udvd_nat_alt: \a udvd b \ (\n. unat b = n * unat a)\ by (auto simp add: udvd_iff_dvd) lemma udvd_unfold_int: \a udvd b \ (\n\0. uint b = n * uint a)\ unfolding udvd_iff_dvd_int by (metis dvd_div_mult_self dvd_triv_right uint_div_distrib uint_ge_0) lemma unat_minus_one: \unat (w - 1) = unat w - 1\ if \w \ 0\ proof - have "0 \ uint w" by (fact uint_nonnegative) moreover from that have "0 \ uint w" by (simp add: uint_0_iff) ultimately have "1 \ uint w" by arith from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" by arith with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" by (auto intro: mod_pos_pos_trivial) with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" by (auto simp del: nat_uint_eq) then show ?thesis by (simp only: unat_eq_nat_uint word_arith_wis mod_diff_right_eq) (metis of_int_1 uint_word_of_int unsigned_1) qed lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" for x :: "'a::len word" and y :: "'b::len word" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection \Conditions for the addition (etc) of two words to overflow\ lemma uint_add_lem: "(uint x + uint y < 2 ^ LENGTH('a)) = (uint (x + y) = uint x + uint y)" for x y :: "'a::len word" by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) lemma uint_mult_lem: "(uint x * uint y < 2 ^ LENGTH('a)) = (uint (x * y) = uint x * uint y)" for x y :: "'a::len word" by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" unfolding uint_word_ariths by (simp flip: take_bit_eq_mod add: take_bit_int_greater_eq_self_iff) lemma int_mod_ge: \a \ a mod n\ if \a < n\ \0 < n\ for a n :: int -proof (cases \a < 0\) - case True - with \0 < n\ show ?thesis - by (metis less_trans not_less pos_mod_conj) - -next - case False - with \a < n\ show ?thesis - by simp -qed + using that order.trans [of a 0 \a mod n\] by (cases \a < 0\) auto lemma mod_add_if_z: "\x < z; y < z; 0 \ y; 0 \ x; 0 \ z\ \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" for x y z :: int apply (simp add: not_less) by (metis (no_types) add_strict_mono diff_ge_0_iff_ge diff_less_eq minus_mod_self2 mod_pos_pos_trivial) lemma uint_plus_if': "uint (a + b) = (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b else uint a + uint b - 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) lemma mod_sub_if_z: "\x < z; y < z; 0 \ y; 0 \ x; 0 \ z\ \ (x - y) mod z = (if y \ x then x - y else x - y + z)" for x y z :: int using mod_pos_pos_trivial [of "x - y + z" z] by (auto simp add: not_le) lemma uint_sub_if': "uint (a - b) = (if uint b \ uint a then uint a - uint b else uint a - uint b + 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) lemma word_of_int_inverse: "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" for a :: "'a::len word" by transfer (simp add: take_bit_int_eq_self) lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" for x :: "'a::len word" by (auto simp add: unsigned_of_nat take_bit_nat_eq_self) lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" for x :: "'a::len word" by (auto simp add: unsigned_of_nat take_bit_nat_eq_self) lemma un_ui_le: \unat a \ unat b \ uint a \ uint b\ by transfer (simp add: nat_le_iff) lemma unat_plus_if': \unat (a + b) = (if unat a + unat b < 2 ^ LENGTH('a) then unat a + unat b else unat a + unat b - 2 ^ LENGTH('a))\ for a b :: \'a::len word\ apply (auto simp add: not_less le_iff_add) apply (metis (mono_tags, lifting) of_nat_add of_nat_unat take_bit_nat_eq_self_iff unsigned_less unsigned_of_nat unsigned_word_eqI) apply (smt (verit, ccfv_SIG) dbl_simps(3) dbl_simps(5) numerals(1) of_nat_0_le_iff of_nat_add of_nat_eq_iff of_nat_numeral of_nat_power of_nat_unat uint_plus_if' unsigned_1) done lemma unat_sub_if_size: "unat (x - y) = (if unat y \ unat x then unat x - unat y else unat x + 2 ^ size x - unat y)" proof - { assume xy: "\ uint y \ uint x" have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x + 2 ^ LENGTH('a) - uint y)" by simp also have "... = nat (uint x + 2 ^ LENGTH('a)) - nat (uint y)" by (simp add: nat_diff_distrib') also have "... = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" by (metis nat_add_distrib nat_eq_numeral_power_cancel_iff order_less_imp_le unsigned_0 unsigned_greater_eq unsigned_less) finally have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" . } then show ?thesis by (simp add: word_size) (metis nat_diff_distrib' uint_sub_if' un_ui_le unat_eq_nat_uint unsigned_greater_eq) qed lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma uint_split: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" for x :: "'a::len word" by transfer (auto simp add: take_bit_eq_mod) lemma uint_split_asm: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" for x :: "'a::len word" by (auto simp add: unsigned_of_int take_bit_int_eq_self) subsection \Some proof tool support\ \ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ lemma power_False_cong: "False \ a ^ b = c ^ d" by auto lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat_eq_iff unat_sub_if' unat_plus_if' unat_div unat_mod lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint_eq_iff uint_sub_if' uint_plus_if' \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ val unat_arith_simpset = @{context} (* TODO: completely explicitly determined simpset *) |> fold Simplifier.add_simp @{thms unat_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset unat_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms unat_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) \ method_setup unat_arith = \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ "solving word arithmetic via natural numbers and arith" \ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ ML \ val uint_arith_simpset = @{context} (* TODO: completely explicitly determined simpset *) |> fold Simplifier.add_simp @{thms uint_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of; fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset uint_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms uint_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n THEN assume_tac ctxt n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) \ method_setup uint_arith = \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ "solving word arithmetic via integers and arith" subsection \More on overflows and monotonicity\ lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" for x y :: "'a::len word" by (auto simp add: word_size word_le_def uint_add_lem uint_sub_lem) lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" for x y :: "'a::len word" by (auto simp add: word_size word_le_def uint_add_lem uint_sub_lem) lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" for x y :: "'a::len word" by (simp add: ac_simps no_olen_add) lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] lemmas word_sub_le = word_sub_le_iff [THEN iffD2] lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" for x :: "'a::len word" by transfer (simp add: take_bit_decr_eq) lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" for x :: "'a::len word" by transfer (simp add: int_one_le_iff_zero_less less_le) lemma sub_wrap_lt: "x < x - z \ x < z" for x z :: "'a::len word" by (simp add: word_less_def uint_sub_lem) (meson linorder_not_le uint_minus_simple_iff uint_sub_lem word_less_iff_unsigned) lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" for x z :: "'a::len word" by (simp add: le_less sub_wrap_lt ac_simps) lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" for x ab c :: "'a::len word" by uint_arith lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" for x ab c :: "'a::len word" by uint_arith lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" for a b c :: "'a::len word" by uint_arith lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" for a b c :: "'a::len word" by uint_arith lemmas le_plus = le_plus' [rotated] lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" for a b c d :: "'a::len word" by uint_arith lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono: "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" for a b c d :: "'a::len word" by uint_arith lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" for a b c :: "'a::len word" by uint_arith lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" for x y z :: "'a::len word" by uint_arith lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" for x y z :: "'a::len word" by uint_arith lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" for x y z :: "'a::len word" by uint_arith lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" for x z k :: "'a::len word" by uint_arith lemma inc_le: "i < m \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma udvd_incr_lem: "up < uq \ up = ua + n * uint K \ uq = ua + n' * uint K \ up + uint K \ uq" by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) lemma udvd_incr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p + K \ q" unfolding word_less_alt word_le_def by (metis (full_types) order_trans udvd_incr_lem uint_add_le) lemma udvd_decr': assumes "p < q" "uint p = ua + n * uint K" "uint q = ua + n' * uint K" shows "uint q = ua + n' * uint K \ p \ q - K" proof - have "\w wa. uint (w::'a word) \ uint wa + uint (w - wa)" by (metis (no_types) add_diff_cancel_left' diff_add_cancel uint_add_le) moreover have "uint K + uint p \ uint q" using assms by (metis (no_types) add_diff_cancel_left' diff_add_cancel udvd_incr_lem word_less_def) ultimately show ?thesis by (meson add_le_cancel_left order_trans word_less_eq_iff_unsigned) qed lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" unfolding udvd_unfold_int by (meson udvd_decr0) lemma udvd_incr2_K: "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ 0 < K \ p \ p + K \ p + K \ a + s" unfolding udvd_unfold_int apply (simp add: uint_arith_simps split: if_split_asm) apply (metis (no_types, opaque_lifting) le_add_diff_inverse le_less_trans udvd_incr_lem) using uint_lt2p [of s] by simp subsection \Arithmetic type class instantiations\ lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN antisym_conv1] lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" by simp text \ note that \iszero_def\ is only for class \comm_semiring_1_cancel\, which requires word length \\ 1\, ie \'a::len word\ \ lemma iszero_word_no [simp]: "iszero (numeral bin :: 'a::len word) = iszero (take_bit LENGTH('a) (numeral bin :: int))" by (metis iszero_def uint_0_iff uint_bintrunc) text \Use \iszero\ to simplify equalities between word numerals.\ lemmas word_eq_numeral_iff_iszero [simp] = eq_numeral_iff_iszero [where 'a="'a::len word"] subsection \Word and nat\ lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" by (metis of_nat_unat ucast_id unsigned_less) lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" for w :: "'a::len word" using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] by (auto simp flip: take_bit_eq_mod simp add: unsigned_of_nat) lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" by (simp add: of_nat_eq) lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" by (auto simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" by (simp add: wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by transfer (simp add: ac_simps) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by simp lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by simp lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by simp lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemma unat_cong: "x = y \ unat x = unat y" by (fact arg_cong) lemma unat_of_nat: \unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\ by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq) lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat]] lemmas word_sub_less_iff = word_sub_le_iff [unfolded linorder_not_less [symmetric] Not_eq_iff] lemma unat_add_lem: "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" for x y :: "'a::len word" by (metis mod_less unat_word_ariths(1) unsigned_less) lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" by (metis mod_less unat_word_ariths(2) unsigned_less) lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len word" using word_le_plus_either by blast lemma uint_div: \uint (x div y) = uint x div uint y\ by (fact uint_div_distrib) lemma uint_mod: \uint (x mod y) = uint x mod uint y\ by (fact uint_mod_distrib) lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem] lemma word_div_mult: "\0 < y; unat x * unat y < 2 ^ LENGTH('a)\ \ x * y div y = x" for x y :: "'a::len word" by (simp add: unat_eq_zero unat_mult_lem word_arith_nat_div) lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" by unat_arith (meson le_less_trans less_mult_imp_div_less not_le unsigned_less) lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "\i < k div x; 0 < x\ \ i * x < k" for i k x :: "'a::len word" by (metis div_le_mono div_lt'' not_le unat_div word_div_mult word_less_iff_unsigned) lemma div_le_mult: "\i \ k div x; 0 < x\ \ i * x \ k" for i k x :: "'a::len word" by (metis div_lt' less_mult_imp_div_less not_less unat_arith_simps(2) unat_div unat_mult_lem) lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" unfolding uint_nat by (metis div_lt' int_ops(7) of_nat_unat uint_mult_lem unat_mult_lem) lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" for x y z :: "'a::len word" by (metis add.commute diff_add_cancel no_olen_add) lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x lemma le_unat_uoi: \y \ unat z \ unat (word_of_nat y :: 'a word) = y\ for z :: \'a::len word\ by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans) lemmas thd = times_div_less_eq_dividend lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" for n b :: "'a::len word" by (fact div_mult_mod_eq) lemma word_div_mult_le: "a div b * b \ a" for a b :: "'a::len word" by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) lemma word_mod_less_divisor: "0 < n \ m mod n < n" for m n :: "'a::len word" by (simp add: unat_arith_simps) lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" by (induct n) (simp_all add: wi_hom_mult [symmetric]) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" for n :: "'a::len word" by unat_arith subsection \Cardinality, finiteness of set of words\ lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ unfolding inj_on_def by (metis atLeastLessThan_iff word_of_int_inverse) lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ apply transfer apply (auto simp add: image_iff) apply (metis take_bit_int_eq_self_iff) done lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ by (auto simp add: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split) lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" by (simp add: UNIV_eq card_image inj_on_word_of_int) lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len word" unfolding word_size by (rule card_word) end instance word :: (len) finite by standard (simp add: UNIV_eq) subsection \Bitwise Operations on Words\ context includes bit_operations_syntax begin lemma word_wi_log_defs: "NOT (word_of_int a) = word_of_int (NOT a)" "word_of_int a AND word_of_int b = word_of_int (a AND b)" "word_of_int a OR word_of_int b = word_of_int (a OR b)" "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" by (transfer, rule refl)+ lemma word_no_log_defs [simp]: "NOT (numeral a) = word_of_int (NOT (numeral a))" "NOT (- numeral a) = word_of_int (NOT (- numeral a))" "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" by (transfer, rule refl)+ text \Special cases for when one of the arguments equals 1.\ lemma word_bitwise_1_simps [simp]: "NOT (1::'a::len word) = -2" "1 AND numeral b = word_of_int (1 AND numeral b)" "1 AND - numeral b = word_of_int (1 AND - numeral b)" "numeral a AND 1 = word_of_int (numeral a AND 1)" "- numeral a AND 1 = word_of_int (- numeral a AND 1)" "1 OR numeral b = word_of_int (1 OR numeral b)" "1 OR - numeral b = word_of_int (1 OR - numeral b)" "numeral a OR 1 = word_of_int (numeral a OR 1)" "- numeral a OR 1 = word_of_int (- numeral a OR 1)" "1 XOR numeral b = word_of_int (1 XOR numeral b)" "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" "numeral a XOR 1 = word_of_int (numeral a XOR 1)" "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" apply (simp_all add: word_uint_eq_iff unsigned_not_eq unsigned_and_eq unsigned_or_eq unsigned_xor_eq of_nat_take_bit ac_simps unsigned_of_int) apply (simp_all add: minus_numeral_eq_not_sub_one) apply (simp_all only: sub_one_eq_not_neg bit.xor_compl_right take_bit_xor bit.double_compl) apply simp_all done text \Special cases for when one of the arguments equals -1.\ lemma word_bitwise_m1_simps [simp]: "NOT (-1::'a::len word) = 0" "(-1::'a::len word) AND x = x" "x AND (-1::'a::len word) = x" "(-1::'a::len word) OR x = -1" "x OR (-1::'a::len word) = -1" " (-1::'a::len word) XOR x = NOT x" "x XOR (-1::'a::len word) = NOT x" by (transfer, simp)+ lemma word_of_int_not_numeral_eq [simp]: \(word_of_int (NOT (numeral bin)) :: 'a::len word) = - numeral bin - 1\ by transfer (simp add: not_eq_complement) lemma uint_and: \uint (x AND y) = uint x AND uint y\ by transfer simp lemma uint_or: \uint (x OR y) = uint x OR uint y\ by transfer simp lemma uint_xor: \uint (x XOR y) = uint x XOR uint y\ by transfer simp \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add word_wi_log_defs lemma word_bw_assocs: "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" for x :: "'a::len word" by (fact ac_simps)+ lemma word_bw_comms: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: "'a::len word" by (fact ac_simps)+ lemma word_bw_lcs: "y AND x AND z = x AND y AND z" "y OR x OR z = x OR y OR z" "y XOR x XOR z = x XOR y XOR z" for x :: "'a::len word" by (fact ac_simps)+ lemma word_log_esimps: "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" for x :: "'a::len word" by simp_all lemma word_not_dist: "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" for x :: "'a::len word" by simp_all lemma word_bw_same: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: "'a::len word" by simp_all lemma word_ao_absorbs [simp]: "x AND (y OR x) = x" "x OR y AND x = x" "x AND (x OR y) = x" "y AND x OR x = x" "(y OR x) AND x = x" "x OR x AND y = x" "(x OR y) AND x = x" "x AND y OR x = x" for x :: "'a::len word" by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" by (fact bit.double_compl) lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len word" by (fact bit.conj_disj_distrib2) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len word" by (fact bit.disj_conj_distrib2) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" by (simp add: not_eq_complement) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len word" by transfer (simp add: plus_and_or) lemma leoa: "w = x OR y \ y = w AND y" for x :: "'a::len word" by auto lemma leao: "w' = x' AND y' \ x' = x' OR w'" for x' :: "'a::len word" by auto lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" for w w' :: "'a::len word" by (auto intro: leoa leao) lemma le_word_or2: "x \ x OR y" for x y :: "'a::len word" by (simp add: or_greater_eq uint_or word_le_def) lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] lemma bit_horner_sum_bit_word_iff [bit_simps]: \bit (horner_sum of_bool (2 :: 'a::len word) bs) n \ n < min LENGTH('a) (length bs) \ bs ! n\ by transfer (simp add: bit_horner_sum_bit_iff) definition word_reverse :: \'a::len word \ 'a word\ where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. lemma bit_word_reverse_iff [bit_simps]: \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ for w :: \'a::len word\ by (cases \n < LENGTH('a)\) (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" by (cases \n < LENGTH('a)\; transfer; force) lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) subsubsection \shift functions in terms of lists of bools\ lemma drop_bit_word_numeral [simp]: \drop_bit (numeral n) (numeral k) = (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer simp lemma drop_bit_word_Suc_numeral [simp]: \drop_bit (Suc n) (numeral k) = (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer simp lemma drop_bit_word_minus_numeral [simp]: \drop_bit (numeral n) (- numeral k) = (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\ by transfer simp lemma drop_bit_word_Suc_minus_numeral [simp]: \drop_bit (Suc n) (- numeral k) = (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\ by transfer simp lemma signed_drop_bit_word_numeral [simp]: \signed_drop_bit (numeral n) (numeral k) = (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\ by transfer simp lemma signed_drop_bit_word_Suc_numeral [simp]: \signed_drop_bit (Suc n) (numeral k) = (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\ by transfer simp lemma signed_drop_bit_word_minus_numeral [simp]: \signed_drop_bit (numeral n) (- numeral k) = (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\ by transfer simp lemma signed_drop_bit_word_Suc_minus_numeral [simp]: \signed_drop_bit (Suc n) (- numeral k) = (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\ by transfer simp lemma take_bit_word_numeral [simp]: \take_bit (numeral n) (numeral k) = (word_of_int (take_bit (min LENGTH('a) (numeral n)) (numeral k)) :: 'a::len word)\ by transfer rule lemma take_bit_word_Suc_numeral [simp]: \take_bit (Suc n) (numeral k) = (word_of_int (take_bit (min LENGTH('a) (Suc n)) (numeral k)) :: 'a::len word)\ by transfer rule lemma take_bit_word_minus_numeral [simp]: \take_bit (numeral n) (- numeral k) = (word_of_int (take_bit (min LENGTH('a) (numeral n)) (- numeral k)) :: 'a::len word)\ by transfer rule lemma take_bit_word_Suc_minus_numeral [simp]: \take_bit (Suc n) (- numeral k) = (word_of_int (take_bit (min LENGTH('a) (Suc n)) (- numeral k)) :: 'a::len word)\ by transfer rule lemma signed_take_bit_word_numeral [simp]: \signed_take_bit (numeral n) (numeral k) = (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer rule lemma signed_take_bit_word_Suc_numeral [simp]: \signed_take_bit (Suc n) (numeral k) = (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer rule lemma signed_take_bit_word_minus_numeral [simp]: \signed_take_bit (numeral n) (- numeral k) = (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\ by transfer rule lemma signed_take_bit_word_Suc_minus_numeral [simp]: \signed_take_bit (Suc n) (- numeral k) = (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\ by transfer rule lemma False_map2_or: "\set xs \ {False}; length ys = length xs\ \ map2 (\) xs ys = ys" by (induction xs arbitrary: ys) (auto simp: length_Suc_conv) lemma align_lem_or: assumes "length xs = n + m" "length ys = n + m" and "drop m xs = replicate n False" "take m ys = replicate m False" shows "map2 (\) xs ys = take m xs @ drop m ys" using assms proof (induction xs arbitrary: ys m) case (Cons a xs) then show ?case by (cases m) (auto simp: length_Suc_conv False_map2_or) qed auto lemma False_map2_and: "\set xs \ {False}; length ys = length xs\ \ map2 (\) xs ys = xs" by (induction xs arbitrary: ys) (auto simp: length_Suc_conv) lemma align_lem_and: assumes "length xs = n + m" "length ys = n + m" and "drop m xs = replicate n False" "take m ys = replicate m False" shows "map2 (\) xs ys = replicate (n + m) False" using assms proof (induction xs arbitrary: ys m) case (Cons a xs) then show ?case by (cases m) (auto simp: length_Suc_conv set_replicate_conv_if False_map2_and) qed auto subsubsection \Mask\ lemma minus_1_eq_mask: \- 1 = (mask LENGTH('a) :: 'a::len word)\ by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff) lemma mask_eq_decr_exp: \mask n = 2 ^ n - (1 :: 'a::len word)\ by (fact mask_eq_exp_minus_1) lemma mask_Suc_rec: \mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1) context begin qualified lemma bit_mask_iff [bit_simps]: \bit (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ by (simp add: bit_mask_iff not_le) end lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))" by transfer simp lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" by transfer (simp add: ac_simps take_bit_eq_mask) lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" by (simp add: take_bit_eq_mask of_int_and_eq of_int_mask_eq) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" by (auto simp add: and_mask_wi min_def wi_bintr) lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" by (simp only: and_mask_bintr take_bit_eq_mod) lemma uint_mask_eq: \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" by (metis take_bit_eq_mask take_bit_int_less_exp unsigned_take_bit_eq) lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (auto simp flip: take_bit_eq_mask) apply (metis take_bit_int_eq_self_iff uint_take_bit_eq) apply (simp add: take_bit_int_eq_self unsigned_take_bit_eq word_uint_eqI) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" for w :: "'a::len word" by transfer simp lemma less_mask_eq: fixes x :: "'a::len word" assumes "x < 2 ^ n" shows "x AND mask n = x" by (metis (no_types) assms lt2p_lem mask_eq_iff not_less word_2p_lem word_size) lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" for x :: \'a::len word\ unfolding word_size by (erule and_mask_less') lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" for c x :: "'a::len word" by (auto simp: word_mod_def uint_2p and_mask_mod_2p) lemma mask_eqs: "(a AND mask n) + b AND mask n = a + b AND mask n" "a + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - b AND mask n = a - b AND mask n" "a - (b AND mask n) AND mask n = a - b AND mask n" "a * (b AND mask n) AND mask n = a * b AND mask n" "(b AND mask n) * a AND mask n = b * a AND mask n" "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" "- (a AND mask n) AND mask n = - a AND mask n" "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] unfolding take_bit_eq_mask [symmetric] by (transfer; simp add: take_bit_eq_mod mod_simps)+ lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" for x :: \'a::len word\ using word_of_int_Ex [where x=x] unfolding take_bit_eq_mask [symmetric] by (transfer; simp add: take_bit_eq_mod mod_simps)+ lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by transfer simp subsubsection \Slices\ definition slice1 :: \nat \ 'a::len word \ 'b::len word\ where \slice1 n w = (if n < LENGTH('a) then ucast (drop_bit (LENGTH('a) - n) w) else push_bit (n - LENGTH('a)) (ucast w))\ lemma bit_slice1_iff [bit_simps]: \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ for w :: \'a::len word\ by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff not_less not_le ac_simps dest: bit_imp_le_length) definition slice :: \nat \ 'a::len word \ 'b::len word\ where \slice n = slice1 (LENGTH('a) - n)\ lemma bit_slice_iff [bit_simps]: \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ for w :: \'a::len word\ by (simp add: slice_def word_size bit_slice1_iff) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp lemma slice_0 [simp] : "slice n 0 = 0" unfolding slice_def by auto lemma ucast_slice1: "ucast w = slice1 (size w) w" unfolding slice1_def by (simp add: size_word.rep_eq) lemma ucast_slice: "ucast w = slice 0 w" by (simp add: slice_def slice1_def) lemma slice_id: "slice 0 t = t" by (simp only: ucast_slice [symmetric] ucast_id) lemma rev_slice1: \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ if \n + k = LENGTH('a) + LENGTH('b)\ proof (rule bit_word_eqI) fix m assume *: \m < LENGTH('a)\ from that have **: \LENGTH('b) = n + k - LENGTH('a)\ by simp show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ unfolding bit_slice1_iff bit_word_reverse_iff using * ** by (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) auto qed lemma rev_slice: "n + k + LENGTH('a::len) = LENGTH('b::len) \ slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" unfolding slice_def word_size by (simp add: rev_slice1) subsubsection \Revcast\ definition revcast :: \'a::len word \ 'b::len word\ where \revcast = slice1 LENGTH('b)\ lemma bit_revcast_iff [bit_simps]: \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ for w :: \'a::len word\ by (simp add: revcast_def bit_slice1_iff) lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" by (simp add: revcast_def word_size) lemma revcast_rev_ucast [OF refl refl refl]: "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ rc = word_reverse uc" by (metis rev_slice1 revcast_slice1 ucast_slice1 word_size) lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" using revcast_rev_ucast [of "word_reverse w"] by simp lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" by (fact revcast_rev_ucast [THEN word_rev_gal']) lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" by (fact revcast_ucast [THEN word_rev_gal']) text "linking revcast and cast via shift" lemmas wsst_TYs = source_size target_size word_size lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_eq \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" by transfer (simp add: take_bit_concat_bit_eq) lemma word_cat_split_alt: "\size w \ size u + size v; word_split w = (u,v)\ \ word_cat u v = w" unfolding word_split_def by (rule bit_word_eqI) (auto simp add: bit_word_cat_iff not_less word_size bit_ucast_iff bit_drop_bit_eq) lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: assumes "word_split w = (u, v)" shows "u = slice (size v) w \ v = slice 0 w" unfolding word_size proof (intro conjI) have \
: "\n. \ucast (drop_bit LENGTH('b) w) = u; LENGTH('c) < LENGTH('b)\ \ \ bit u n" by (metis bit_take_bit_iff bit_word_of_int_iff diff_is_0_eq' drop_bit_take_bit less_imp_le less_nat_zero_code of_int_uint unsigned_drop_bit_eq) show "u = slice LENGTH('b) w" proof (rule bit_word_eqI) show "bit u n = bit ((slice LENGTH('b) w)::'a word) n" if "n < LENGTH('a)" for n using assms bit_imp_le_length unfolding word_split_def bit_slice_iff by (fastforce simp add: \
ac_simps word_size bit_ucast_iff bit_drop_bit_eq) qed show "v = slice 0 w" by (metis Pair_inject assms ucast_slice word_split_bin') qed lemma slice_cat1 [OF refl]: "\wc = word_cat a b; size a + size b \ size wc\ \ slice (size b) wc = a" by (rule bit_word_eqI) (auto simp add: bit_slice_iff bit_word_cat_iff word_size) lemmas slice_cat2 = trans [OF slice_id word_cat_id] lemma cat_slices: "\a = slice n c; b = slice 0 c; n = size b; size c \ size a + size b\ \ word_cat a b = c" by (rule bit_word_eqI) (auto simp add: bit_slice_iff bit_word_cat_iff word_size) lemma word_split_cat_alt: assumes "w = word_cat u v" and size: "size u + size v \ size w" shows "word_split w = (u,v)" proof - have "ucast ((drop_bit LENGTH('c) (word_cat u v))::'a word) = u" "ucast ((word_cat u v)::'a word) = v" using assms by (auto simp add: word_size bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff intro: bit_eqI) then show ?thesis by (simp add: assms(1) word_split_bin') qed lemma horner_sum_uint_exp_Cons_eq: \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ for ws :: \'a::len word list\ by (simp add: bintr_uint concat_bit_eq push_bit_eq_mult) lemma bit_horner_sum_uint_exp_iff: \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ for ws :: \'a::len word list\ proof (induction ws arbitrary: n) case Nil then show ?case by simp next case (Cons w ws) then show ?case by (cases \n \ LENGTH('a)\) (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) qed subsection \Rotation\ lemma word_rotr_word_rotr_eq: \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) lemma word_rot_lem: "\l + k = d + k mod l; n < l\ \ ((d + n) mod l) = n" for l::nat by (metis (no_types, lifting) add.commute add.right_neutral add_diff_cancel_left' mod_if mod_mult_div_eq mod_mult_self2 mod_self) lemma word_rot_rl [simp]: \word_rotl k (word_rotr k v) = v\ proof (rule bit_word_eqI) show "bit (word_rotl k (word_rotr k v)) n = bit v n" if "n < LENGTH('a)" for n using that by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split) qed lemma word_rot_lr [simp]: \word_rotr k (word_rotl k v) = v\ proof (rule bit_word_eqI) show "bit (word_rotr k (word_rotl k v)) n = bit v n" if "n < LENGTH('a)" for n using that by (auto simp add: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split) qed lemma word_rot_gal: \word_rotr n v = w \ word_rotl n w = v\ by auto lemma word_rot_gal': \w = word_rotr n v \ v = word_rotl n w\ by auto lemma word_rotr_rev: \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('a)\ moreover have \1 + ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = int LENGTH('a)\ apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod int LENGTH('a) = 0\) using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] apply simp_all apply (auto simp add: algebra_simps) apply (metis (mono_tags, opaque_lifting) Abs_fnat_hom_add mod_Suc mod_mult_self2_is_0 of_nat_Suc of_nat_mod semiring_char_0_class.of_nat_neq_0) apply (metis (no_types, opaque_lifting) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) done then have \int ((m + n) mod LENGTH('a)) = int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ using \m < LENGTH('a)\ by (simp only: of_nat_mod mod_simps) (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) then have \(m + n) mod LENGTH('a) = LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ by simp ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) qed lemma word_roti_0 [simp]: "word_roti 0 w = w" by transfer simp lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" by (rule bit_word_eqI) (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps) lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w" by transfer simp lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] end subsubsection \"Word rotation commutes with bit-wise operations\ \ \using locale to not pollute lemma namespace\ locale word_rotate begin context includes bit_operations_syntax begin lemma word_rot_logs: "word_rotl n (NOT v) = NOT (word_rotl n v)" "word_rotr n (NOT v) = NOT (word_rotr n v)" "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" by (rule bit_word_eqI, auto simp add: bit_word_rotl_iff bit_word_rotr_iff bit_and_iff bit_or_iff bit_xor_iff bit_not_iff algebra_simps not_le)+ end end lemmas word_rot_logs = word_rotate.word_rot_logs lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" by transfer simp_all lemma word_roti_0' [simp] : "word_roti n 0 = 0" by transfer simp declare word_roti_eq_word_rotr_word_rotl [simp] subsection \Maximum machine word\ context includes bit_operations_syntax begin lemma word_int_cases: fixes x :: "'a::len word" obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" by (rule that [of \uint x\]) simp_all lemma word_nat_cases [cases type: word]: fixes x :: "'a::len word" obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" by (rule that [of \unat x\]) simp_all lemma max_word_max [intro!]: \n \ - 1\ for n :: \'a::len word\ by (fact word_order.extremum) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" by simp lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" by (fact word_exp_length_eq_0) lemma max_word_wrap: \x + 1 = 0 \ x = - 1\ for x :: \'a::len word\ by (simp add: eq_neg_iff_add_eq_0) lemma word_and_max: \x AND - 1 = x\ for x :: \'a::len word\ by (fact word_log_esimps) lemma word_or_max: \x OR - 1 = - 1\ for x :: \'a::len word\ by (fact word_log_esimps) lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" for x y z :: "'a::len word" by (fact bit.conj_disj_distrib) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len word" by (fact bit.disj_conj_distrib) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len word" by (fact bit.conj_cancel_right) lemma word_or_not [simp]: \x OR NOT x = - 1\ for x :: \'a::len word\ by (fact bit.disj_cancel_right) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len word" by (fact bit.xor_def) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma word_less_1 [simp]: "x < 1 \ x = 0" for x :: "'a::len word" by (simp add: word_less_nat_alt unat_0_iff) lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" by (simp add: take_bit_eq_mod word_size uint_word_of_int_eq uint_plus_if') lemma unat_plus_if_size: "unat (x + y) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" for x y :: "'a::len word" by (simp add: size_word.rep_eq unat_arith_simps) lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" by (fact word_coorder.not_eq_extremum) lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" for c :: "'a::len word" by (fact unat_div) lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" by (simp add: size_word.rep_eq uint_sub_if') lemma unat_sub: \unat (a - b) = unat a - unat b\ if \b \ a\ by (meson that unat_sub_if_size word_le_nat_alt) lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" by simp lemma word_of_int_inj: \(word_of_int x :: 'a::len word) = word_of_int y \ x = y\ if \0 \ x \ x < 2 ^ LENGTH('a)\ \0 \ y \ y < 2 ^ LENGTH('a)\ using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) lemma word_le_less_eq: "x \ y \ x = y \ x < y" for x y :: "'z::len word" by (auto simp add: order_class.le_less) lemma mod_plus_cong: fixes b b' :: int assumes 1: "b = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: fixes b b' :: int assumes "b = b'" and "x mod b' = x' mod b'" and "y mod b' = y' mod b'" and "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms [symmetric] by (auto intro: mod_diff_cong) lemma word_induct_less [case_names zero less]: \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ for m :: \'a::len word\ proof - define q where \q = unat m\ with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ by simp then have \P (word_of_nat q :: 'a word)\ proof (induction q) case 0 show ?case by (simp add: zero) next case (Suc q) show ?case proof (cases \1 + word_of_nat q = (0 :: 'a word)\) case True then show ?thesis by (simp add: zero) next case False then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ by (simp add: unatSuc word_less_nat_alt) then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) have \P (word_of_nat q)\ by (simp add: "**" Suc.IH Suc.prems) with * have \P (1 + word_of_nat q)\ by (rule Suc.prems) then show ?thesis by simp qed qed with \q = unat m\ show ?thesis by simp qed lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" by (rule word_induct_less) lemma word_induct2 [case_names zero suc, induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" by (induction rule: word_induct_less; force) subsection \Recursion combinator for words\ definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0 [simp]: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc [simp]: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" by (simp add: unatSuc word_rec_def) lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" by (metis add.commute diff_add_cancel word_rec_Suc) lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) simp_all lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" by (induct n) simp_all lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" proof (induction n arbitrary: z f) case zero then show ?case by (metis diff_0_right word_le_0_iff word_rec_0) next case (suc n z f) show ?case proof (cases "1 + (n - m) = 0") case True then show ?thesis by (simp add: add_diff_eq) next case False then have eq: "1 + n - m = 1 + (n - m)" by simp with False have "m \ n" by (metis "suc.prems" add.commute dual_order.antisym eq_iff_diff_eq_0 inc_le leI) with False "suc.hyps" show ?thesis using suc.IH [of "f 0 z" "f \ (+) 1"] by (simp add: word_rec_in2 eq add.assoc o_def) qed qed lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) auto lemma word_rec_id_eq: "(\m. m < n \ f m = id) \ word_rec z f n = z" by (induction n) (auto simp add: unatSuc unat_arith_simps(2)) lemma word_rec_max: assumes "\m\n. m \ - 1 \ f m = id" shows "word_rec z f (- 1) = word_rec z f n" proof - have \
: "\m. \m < - 1 - n\ \ (f \ (+) n) m = id" using assms by (metis (mono_tags, lifting) add.commute add_diff_cancel_left' comp_apply less_le olen_add_eqv plus_minus_no_overflow word_n1_ge) have "word_rec z f (- 1) = word_rec (word_rec z f (- 1 - (- 1 - n))) (f \ (+) (- 1 - (- 1 - n))) (- 1 - n)" by (meson word_n1_ge word_rec_twice) also have "... = word_rec z f n" by (metis (no_types, lifting) \
diff_add_cancel minus_diff_eq uminus_add_conv_diff word_rec_id_eq) finally show ?thesis . qed end subsection \Tool support\ ML_file \Tools/smt_word.ML\ end diff --git a/src/HOL/Number_Theory/Cong.thy b/src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy +++ b/src/HOL/Number_Theory/Cong.thy @@ -1,754 +1,750 @@ (* Title: HOL/Number_Theory/Cong.thy Author: Christophe Tabacznyj Author: Lawrence C. Paulson Author: Amine Chaieb Author: Thomas M. Rasmussen Author: Jeremy Avigad Defines congruence (notation: [x = y] (mod z)) for natural numbers and integers. This file combines and revises a number of prior developments. The original theories "GCD" and "Primes" were by Christophe Tabacznyj and Lawrence C. Paulson, based on @{cite davenport92}. They introduced gcd, lcm, and prime for the natural numbers. The original theory "IntPrimes" was by Thomas M. Rasmussen, and extended gcd, lcm, primes to the integers. Amine Chaieb provided another extension of the notions to the integers, and added a number of results to "Primes" and "GCD". The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and developed the congruence relations on the integers. The notion was extended to the natural numbers by Chaieb. Jeremy Avigad combined these, revised and tidied them, made the development uniform for the natural numbers and the integers, and added a number of new theorems. *) section \Congruence\ theory Cong imports "HOL-Computational_Algebra.Primes" begin subsection \Generic congruences\ context unique_euclidean_semiring begin definition cong :: "'a \ 'a \ 'a \ bool" (\(1[_ = _] '(' mod _'))\) where "cong b c a \ b mod a = c mod a" abbreviation notcong :: "'a \ 'a \ 'a \ bool" (\(1[_ \ _] '(' mod _'))\) where "notcong b c a \ \ cong b c a" lemma cong_refl [simp]: "[b = b] (mod a)" by (simp add: cong_def) lemma cong_sym: "[b = c] (mod a) \ [c = b] (mod a)" by (simp add: cong_def) lemma cong_sym_eq: "[b = c] (mod a) \ [c = b] (mod a)" by (auto simp add: cong_def) lemma cong_trans [trans]: "[b = c] (mod a) \ [c = d] (mod a) \ [b = d] (mod a)" by (simp add: cong_def) lemma cong_mult_self_right: "[b * a = 0] (mod a)" by (simp add: cong_def) lemma cong_mult_self_left: "[a * b = 0] (mod a)" by (simp add: cong_def) lemma cong_mod_left [simp]: "[b mod a = c] (mod a) \ [b = c] (mod a)" by (simp add: cong_def) lemma cong_mod_right [simp]: "[b = c mod a] (mod a) \ [b = c] (mod a)" by (simp add: cong_def) lemma cong_0 [simp, presburger]: "[b = c] (mod 0) \ b = c" by (simp add: cong_def) lemma cong_1 [simp, presburger]: "[b = c] (mod 1)" by (simp add: cong_def) lemma cong_dvd_iff: "a dvd b \ a dvd c" if "[b = c] (mod a)" using that by (auto simp: cong_def dvd_eq_mod_eq_0) lemma cong_0_iff: "[b = 0] (mod a) \ a dvd b" by (simp add: cong_def dvd_eq_mod_eq_0) lemma cong_add: "[b = c] (mod a) \ [d = e] (mod a) \ [b + d = c + e] (mod a)" by (auto simp add: cong_def intro: mod_add_cong) lemma cong_mult: "[b = c] (mod a) \ [d = e] (mod a) \ [b * d = c * e] (mod a)" by (auto simp add: cong_def intro: mod_mult_cong) lemma cong_scalar_right: "[b = c] (mod a) \ [b * d = c * d] (mod a)" by (simp add: cong_mult) lemma cong_scalar_left: "[b = c] (mod a) \ [d * b = d * c] (mod a)" by (simp add: cong_mult) lemma cong_pow: "[b = c] (mod a) \ [b ^ n = c ^ n] (mod a)" by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a]) lemma cong_sum: "[sum f A = sum g A] (mod a)" if "\x. x \ A \ [f x = g x] (mod a)" using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add) lemma cong_prod: "[prod f A = prod g A] (mod a)" if "(\x. x \ A \ [f x = g x] (mod a))" using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult) lemma mod_mult_cong_right: "[c mod (a * b) = d] (mod a) \ [c = d] (mod a)" by (simp add: cong_def mod_mod_cancel mod_add_left_eq) lemma mod_mult_cong_left: "[c mod (b * a) = d] (mod a) \ [c = d] (mod a)" using mod_mult_cong_right [of c a b d] by (simp add: ac_simps) end context unique_euclidean_ring begin lemma cong_diff: "[b = c] (mod a) \ [d = e] (mod a) \ [b - d = c - e] (mod a)" by (auto simp add: cong_def intro: mod_diff_cong) lemma cong_diff_iff_cong_0: "[b - c = 0] (mod a) \ [b = c] (mod a)" (is "?P \ ?Q") proof assume ?P then have "[b - c + c = 0 + c] (mod a)" by (rule cong_add) simp then show ?Q by simp next assume ?Q with cong_diff [of b c a c c] show ?P by simp qed lemma cong_minus_minus_iff: "[- b = - c] (mod a) \ [b = c] (mod a)" using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a] by (simp add: cong_0_iff dvd_diff_commute) lemma cong_modulus_minus_iff [iff]: "[b = c] (mod - a) \ [b = c] (mod a)" using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"] by (simp add: cong_0_iff) lemma cong_iff_dvd_diff: "[a = b] (mod m) \ m dvd (a - b)" by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0) lemma cong_iff_lin: "[a = b] (mod m) \ (\k. b = a + m * k)" (is "?P \ ?Q") proof - have "?P \ m dvd b - a" by (simp add: cong_iff_dvd_diff dvd_diff_commute) also have "\ \ ?Q" by (auto simp add: algebra_simps elim!: dvdE) finally show ?thesis by simp qed lemma cong_add_lcancel: "[a + x = a + y] (mod n) \ [x = y] (mod n)" by (simp add: cong_iff_lin algebra_simps) lemma cong_add_rcancel: "[x + a = y + a] (mod n) \ [x = y] (mod n)" by (simp add: cong_iff_lin algebra_simps) lemma cong_add_lcancel_0: "[a + x = a] (mod n) \ [x = 0] (mod n)" using cong_add_lcancel [of a x 0 n] by simp lemma cong_add_rcancel_0: "[x + a = a] (mod n) \ [x = 0] (mod n)" using cong_add_rcancel [of x a 0 n] by simp lemma cong_dvd_modulus: "[x = y] (mod n)" if "[x = y] (mod m)" and "n dvd m" using that by (auto intro: dvd_trans simp add: cong_iff_dvd_diff) lemma cong_modulus_mult: "[x = y] (mod m)" if "[x = y] (mod m * n)" using that by (simp add: cong_iff_dvd_diff) (rule dvd_mult_left) end lemma cong_abs [simp]: "[x = y] (mod \m\) \ [x = y] (mod m)" for x y :: "'a :: {unique_euclidean_ring, linordered_idom}" by (simp add: cong_iff_dvd_diff) lemma cong_square: "prime p \ 0 < a \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = - 1] (mod p)" for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}" by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest: prime_dvd_multD) lemma cong_mult_rcancel: "[a * k = b * k] (mod m) \ [a = b] (mod m)" if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" using that by (auto simp add: cong_iff_dvd_diff left_diff_distrib [symmetric] ac_simps coprime_dvd_mult_right_iff) lemma cong_mult_lcancel: "[k * a = k * b] (mod m) = [a = b] (mod m)" if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" using that cong_mult_rcancel [of k m a b] by (simp add: ac_simps) lemma coprime_cong_mult: "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" for a b :: "'a :: {unique_euclidean_ring, semiring_gcd}" by (simp add: cong_iff_dvd_diff divides_mult) lemma cong_gcd_eq: "gcd a m = gcd b m" if "[a = b] (mod m)" for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" proof (cases "m = 0") case True with that show ?thesis by simp next case False moreover have "gcd (a mod m) m = gcd (b mod m) m" using that by (simp add: cong_def) ultimately show ?thesis by simp qed lemma cong_imp_coprime: "[a = b] (mod m) \ coprime a m \ coprime b m" for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq) lemma cong_cong_prod_coprime: "[x = y] (mod (\i\A. m i))" if "(\i\A. [x = y] (mod m i))" "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" for x y :: "'a :: {unique_euclidean_ring, semiring_gcd}" using that by (induct A rule: infinite_finite_induct) (auto intro!: coprime_cong_mult prod_coprime_right) subsection \Congruences on \<^typ>\nat\ and \<^typ>\int\\ lemma cong_int_iff: "[int m = int q] (mod int n) \ [m = q] (mod n)" by (simp add: cong_def of_nat_mod [symmetric]) lemma cong_Suc_0 [simp, presburger]: "[m = n] (mod Suc 0)" using cong_1 [of m n] by simp lemma cong_diff_nat: "[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)" and "a \ c" "b \ d" for a b c d m :: nat proof - have "[c + (a - c) = d + (b - d)] (mod m)" using that by simp with \[c = d] (mod m)\ have "[c + (a - c) = c + (b - d)] (mod m)" using mod_add_cong by (auto simp add: cong_def) fastforce then show ?thesis by (simp add: cong_def nat_mod_eq_iff) qed lemma cong_diff_iff_cong_0_nat: "[a - b = 0] (mod m) \ [a = b] (mod m)" if "a \ b" for a b :: nat using that by (simp add: cong_0_iff) (simp add: cong_def mod_eq_dvd_iff_nat) lemma cong_diff_iff_cong_0_nat': "[nat \int a - int b\ = 0] (mod m) \ [a = b] (mod m)" proof (cases "b \ a") case True then show ?thesis by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m]) next case False then have "a \ b" by simp then show ?thesis by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m]) (auto simp add: cong_def) qed lemma cong_altdef_nat: "a \ b \ [a = b] (mod m) \ m dvd (a - b)" for a b :: nat by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat) lemma cong_altdef_nat': "[a = b] (mod m) \ m dvd nat \int a - int b\" using cong_diff_iff_cong_0_nat' [of a b m] by (simp only: cong_0_iff [symmetric]) lemma cong_mult_rcancel_nat: "[a * k = b * k] (mod m) \ [a = b] (mod m)" if "coprime k m" for a k m :: nat proof - have "[a * k = b * k] (mod m) \ m dvd nat \int (a * k) - int (b * k)\" by (simp add: cong_altdef_nat') also have "\ \ m dvd nat \(int a - int b) * int k\" by (simp add: algebra_simps) also have "\ \ m dvd nat \int a - int b\ * k" by (simp add: abs_mult nat_times_as_int) also have "\ \ m dvd nat \int a - int b\" by (rule coprime_dvd_mult_left_iff) (use \coprime k m\ in \simp add: ac_simps\) also have "\ \ [a = b] (mod m)" by (simp add: cong_altdef_nat') finally show ?thesis . qed lemma cong_mult_lcancel_nat: "[k * a = k * b] (mod m) = [a = b] (mod m)" if "coprime k m" for a k m :: nat using that by (simp add: cong_mult_rcancel_nat ac_simps) lemma coprime_cong_mult_nat: "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" for a b :: nat by (simp add: cong_altdef_nat' divides_mult) lemma cong_less_imp_eq_nat: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: nat by (auto simp add: cong_def) lemma cong_less_imp_eq_int: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: int by (auto simp add: cong_def) lemma cong_less_unique_nat: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: nat - by (auto simp: cong_def) (metis mod_less_divisor mod_mod_trivial) + by (auto simp: cong_def) (metis mod_mod_trivial mod_less_divisor) lemma cong_less_unique_int: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: int - by (auto simp: cong_def) (metis mod_mod_trivial pos_mod_conj) + by (auto simp add: cong_def) (metis mod_mod_trivial pos_mod_bound pos_mod_sign) lemma cong_iff_lin_nat: "[a = b] (mod m) \ (\k1 k2. b + k1 * m = a + k2 * m)" for a b :: nat apply (auto simp add: cong_def nat_mod_eq_iff) apply (metis mult.commute) apply (metis mult.commute) done lemma cong_cong_mod_nat: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: nat by simp lemma cong_cong_mod_int: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: int by simp lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \ [x = y] (mod n)" for a x y :: nat by (simp add: cong_iff_lin_nat) lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \ [x = y] (mod n)" for a x y :: nat by (simp add: cong_iff_lin_nat) lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \ [x = 0] (mod n)" for a x :: nat using cong_add_lcancel_nat [of a x 0 n] by simp lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \ [x = 0] (mod n)" for a x :: nat using cong_add_rcancel_nat [of x a 0 n] by simp lemma cong_dvd_modulus_nat: "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" for x y :: nat by (auto simp add: cong_altdef_nat') lemma cong_to_1_nat: fixes a :: nat assumes "[a = 1] (mod n)" shows "n dvd (a - 1)" proof (cases "a = 0") case True then show ?thesis by force next case False with assms show ?thesis by (metis cong_altdef_nat leI less_one) qed lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \ n = Suc 0" by (auto simp: cong_def) lemma cong_0_1_nat: "[0 = 1] (mod n) \ n = 1" for n :: nat by (auto simp: cong_def) lemma cong_0_1_int: "[0 = 1] (mod n) \ n = 1 \ n = - 1" for n :: int by (auto simp: cong_def zmult_eq_1_iff) lemma cong_to_1'_nat: "[a = 1] (mod n) \ a = 0 \ n = 1 \ (\m. a = 1 + m * n)" for a :: nat by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) lemma cong_le_nat: "y \ x \ [x = y] (mod n) \ (\q. x = q * n + y)" for x y :: nat by (auto simp add: cong_altdef_nat le_imp_diff_is_add) lemma cong_solve_nat: fixes a :: nat shows "\x. [a * x = gcd a n] (mod n)" proof (cases "a = 0 \ n = 0") case True then show ?thesis by (force simp add: cong_0_iff cong_sym) next case False then show ?thesis using bezout_nat [of a n] by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) qed lemma cong_solve_int: fixes a :: int shows "\x. [a * x = gcd a n] (mod n)" by (metis bezout_int cong_iff_lin mult.commute) lemma cong_solve_dvd_nat: fixes a :: nat assumes "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - from cong_solve_nat [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast also from assms have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" by auto finally show ?thesis by auto qed lemma cong_solve_dvd_int: fixes a::int assumes b: "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - from cong_solve_int [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast also from b have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" by auto finally show ?thesis by auto qed lemma cong_solve_coprime_nat: "\x. [a * x = Suc 0] (mod n)" if "coprime a n" using that cong_solve_nat [of a n] by auto lemma cong_solve_coprime_int: "\x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int using that cong_solve_int [of a n] by (auto simp add: zabs_def split: if_splits) lemma coprime_iff_invertible_nat: "coprime a m \ (\x. [a * x = Suc 0] (mod m))" (is "?P \ ?Q") proof assume ?P then show ?Q by (auto dest!: cong_solve_coprime_nat) next assume ?Q then obtain b where "[a * b = Suc 0] (mod m)" by blast with coprime_mod_left_iff [of m "a * b"] show ?P by (cases "m = 0 \ m = 1") (unfold cong_def, auto simp add: cong_def) qed lemma coprime_iff_invertible_int: "coprime a m \ (\x. [a * x = 1] (mod m))" (is "?P \ ?Q") for m :: int proof assume ?P then show ?Q by (auto dest: cong_solve_coprime_int) next assume ?Q then obtain b where "[a * b = 1] (mod m)" by blast with coprime_mod_left_iff [of m "a * b"] show ?P by (cases "m = 0 \ m = 1") (unfold cong_def, auto simp add: zmult_eq_1_iff) qed lemma coprime_iff_invertible'_nat: assumes "m > 0" shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = Suc 0] (mod m))" proof - have "\b. \0 < m; [a * b = Suc 0] (mod m)\ \ \b' 0" shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = 1] (mod m))" -proof - - have "\b. \0 < m; [a * b = 1] (mod m)\ \ \b' [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: nat by (meson cong_altdef_nat' lcm_least) lemma cong_cong_lcm_int: "[x = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: int by (auto simp add: cong_iff_dvd_diff lcm_least) lemma cong_cong_prod_coprime_nat: "[x = y] (mod (\i\A. m i))" if "(\i\A. [x = y] (mod m i))" "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" for x y :: nat using that by (induct A rule: infinite_finite_induct) (auto intro!: coprime_cong_mult_nat prod_coprime_right) lemma binary_chinese_remainder_nat: fixes m1 m2 :: nat assumes a: "coprime m1 m2" shows "\x. [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - have "\b1 b2. [b1 = 1] (mod m1) \ [b1 = 0] (mod m2) \ [b2 = 0] (mod m1) \ [b2 = 1] (mod m2)" proof - from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" by (simp add: ac_simps) from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed then obtain b1 b2 where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis by blast qed lemma binary_chinese_remainder_int: fixes m1 m2 :: int assumes a: "coprime m1 m2" shows "\x. [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - have "\b1 b2. [b1 = 1] (mod m1) \ [b1 = 0] (mod m2) \ [b2 = 0] (mod m1) \ [b2 = 1] (mod m2)" proof - from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" by (simp add: ac_simps) from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed then obtain b1 b2 where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis by blast qed lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \ [x = y] (mod m)" for x y :: nat by (metis cong_def mod_mult_cong_right) lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \ x < m \ y < m \ x = y" for x y :: nat by (simp add: cong_def) lemma binary_chinese_remainder_unique_nat: fixes m1 m2 :: nat assumes a: "coprime m1 m2" and nz: "m1 \ 0" "m2 \ 0" shows "\!x. x < m1 * m2 \ [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - obtain y where y1: "[y = u1] (mod m1)" and y2: "[y = u2] (mod m2)" using binary_chinese_remainder_nat [OF a] by blast let ?x = "y mod (m1 * m2)" from nz have less: "?x < m1 * m2" by auto have 1: "[?x = u1] (mod m1)" using y1 mod_mult_cong_right by blast have 2: "[?x = u2] (mod m2)" using y2 mod_mult_cong_left by blast have "z = ?x" if "z < m1 * m2" "[z = u1] (mod m1)" "[z = u2] (mod m2)" for z proof - have "[?x = z] (mod m1)" by (metis "1" cong_def that(2)) moreover have "[?x = z] (mod m2)" by (metis "2" cong_def that(3)) ultimately have "[?x = z] (mod m1 * m2)" using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) with \z < m1 * m2\ \?x < m1 * m2\ show "z = ?x" by (auto simp add: cong_def) qed with less 1 2 show ?thesis by blast qed lemma chinese_remainder_nat: fixes A :: "'a set" and m :: "'a \ nat" and u :: "'a \ nat" assumes fin: "finite A" and cop: "\i \ A. \j \ A. i \ j \ coprime (m i) (m j)" shows "\x. \i \ A. [x = u i] (mod m i)" proof - have "\b. (\i \ A. [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j)))" proof (rule finite_set_choice, rule fin, rule ballI) fix i assume "i \ A" with cop have "coprime (\j \ A - {i}. m j) (m i)" by (intro prod_coprime_left) auto then have "\x. [(\j \ A - {i}. m j) * x = Suc 0] (mod m i)" by (elim cong_solve_coprime_nat) then obtain x where "[(\j \ A - {i}. m j) * x = 1] (mod m i)" by auto moreover have "[(\j \ A - {i}. m j) * x = 0] (mod (\j \ A - {i}. m j))" by (simp add: cong_0_iff) ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod prod m (A - {i}))" by blast qed then obtain b where b: "\i. i \ A \ [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j))" by blast let ?x = "\i\A. (u i) * (b i)" show ?thesis proof (rule exI, clarify) fix i assume a: "i \ A" show "[?x = u i] (mod m i)" proof - from fin a have "?x = (\j \ {i}. u j * b j) + (\j \ A - {i}. u j * b j)" by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong) then have "[?x = u i * b i + (\j \ A - {i}. u j * b j)] (mod m i)" by auto also have "[u i * b i + (\j \ A - {i}. u j * b j) = u i * 1 + (\j \ A - {i}. u j * 0)] (mod m i)" proof (intro cong_add cong_scalar_left cong_sum) show "[b i = 1] (mod m i)" using a b by blast show "[b x = 0] (mod m i)" if "x \ A - {i}" for x proof - have "x \ A" "x \ i" using that by auto then show ?thesis using a b [OF \x \ A\] cong_dvd_modulus_nat fin by blast qed qed finally show ?thesis by simp qed qed qed lemma coprime_cong_prod_nat: "[x = y] (mod (\i\A. m i))" if "\i j. \i \ A; j \ A; i \ j\ \ coprime (m i) (m j)" and "\i. i \ A \ [x = y] (mod m i)" for x y :: nat using that proof (induct A rule: infinite_finite_induct) case (insert x A) then show ?case by simp (metis coprime_cong_mult_nat prod_coprime_right) qed auto lemma chinese_remainder_unique_nat: fixes A :: "'a set" and m :: "'a \ nat" and u :: "'a \ nat" assumes fin: "finite A" and nz: "\i\A. m i \ 0" and cop: "\i\A. \j\A. i \ j \ coprime (m i) (m j)" shows "\!x. x < (\i\A. m i) \ (\i\A. [x = u i] (mod m i))" proof - from chinese_remainder_nat [OF fin cop] obtain y where one: "(\i\A. [y = u i] (mod m i))" by blast let ?x = "y mod (\i\A. m i)" from fin nz have prodnz: "(\i\A. m i) \ 0" by auto then have less: "?x < (\i\A. m i)" by auto have cong: "\i\A. [?x = u i] (mod m i)" using fin one by (auto simp add: cong_def dvd_prod_eqI mod_mod_cancel) have unique: "\z. z < (\i\A. m i) \ (\i\A. [z = u i] (mod m i)) \ z = ?x" proof clarify fix z assume zless: "z < (\i\A. m i)" assume zcong: "(\i\A. [z = u i] (mod m i))" have "\i\A. [?x = z] (mod m i)" using cong zcong by (auto simp add: cong_def) with fin cop have "[?x = z] (mod (\i\A. m i))" by (intro coprime_cong_prod_nat) auto with zless less show "z = ?x" by (auto simp add: cong_def) qed from less cong unique show ?thesis by blast qed end diff --git a/src/HOL/Number_Theory/Pocklington.thy b/src/HOL/Number_Theory/Pocklington.thy --- a/src/HOL/Number_Theory/Pocklington.thy +++ b/src/HOL/Number_Theory/Pocklington.thy @@ -1,1286 +1,1286 @@ (* Title: HOL/Number_Theory/Pocklington.thy Author: Amine Chaieb, Manuel Eberl *) section \Pocklington's Theorem for Primes\ theory Pocklington imports Residues begin subsection \Lemmas about previously defined terms\ lemma prime_nat_iff'': "prime (p::nat) \ p \ 0 \ p \ 1 \ (\m. 0 < m \ m < p \ coprime p m)" apply (auto simp add: prime_nat_iff) apply (rule coprimeI) apply (auto dest: nat_dvd_not_less simp add: ac_simps) apply (metis One_nat_def dvd_1_iff_1 dvd_pos_nat gcd_nat.order_iff is_unit_gcd linorder_neqE_nat nat_dvd_not_less) done lemma finite_number_segment: "card { m. 0 < m \ m < n } = n - 1" proof - have "{ m. 0 < m \ m < n } = {1..Some basic theorems about solving congruences\ lemma cong_solve: fixes n :: nat assumes an: "coprime a n" shows "\x. [a * x = b] (mod n)" proof (cases "a = 0") case True with an show ?thesis by (simp add: cong_def) next case False from bezout_add_strong_nat [OF this] obtain d x y where dxy: "d dvd a" "d dvd n" "a * x = n * y + d" by blast from dxy(1,2) have d1: "d = 1" using assms coprime_common_divisor [of a n d] by simp with dxy(3) have "a * x * b = (n * y + 1) * b" by simp then have "a * (x * b) = n * (y * b) + b" by (auto simp: algebra_simps) then have "a * (x * b) mod n = (n * (y * b) + b) mod n" by simp then have "a * (x * b) mod n = b mod n" by (simp add: mod_add_left_eq) then have "[a * (x * b) = b] (mod n)" by (simp only: cong_def) then show ?thesis by blast qed lemma cong_solve_unique: fixes n :: nat assumes an: "coprime a n" and nz: "n \ 0" shows "\!x. x < n \ [a * x = b] (mod n)" proof - from cong_solve[OF an] obtain x where x: "[a * x = b] (mod n)" by blast let ?P = "\x. x < n \ [a * x = b] (mod n)" let ?x = "x mod n" from x have *: "[a * ?x = b] (mod n)" by (simp add: cong_def mod_mult_right_eq[of a x n]) from mod_less_divisor[ of n x] nz * have Px: "?P ?x" by simp have "y = ?x" if Py: "y < n" "[a * y = b] (mod n)" for y proof - from Py(2) * have "[a * y = a * ?x] (mod n)" by (simp add: cong_def) then have "[y = ?x] (mod n)" by (metis an cong_mult_lcancel_nat) with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz show ?thesis by (simp add: cong_def) qed with Px show ?thesis by blast qed lemma cong_solve_unique_nontrivial: fixes p :: nat assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p" shows "\!y. 0 < y \ y < p \ [x * y = a] (mod p)" proof - from pa have ap: "coprime a p" by (simp add: ac_simps) from x0 xp p have px: "coprime x p" by (auto simp add: prime_nat_iff'' ac_simps) obtain y where y: "y < p" "[x * y = a] (mod p)" "\z. z < p \ [x * z = a] (mod p) \ z = y" by (metis cong_solve_unique neq0_conv p prime_gt_0_nat px) have "y \ 0" proof assume "y = 0" with y(2) have "p dvd a" using cong_dvd_iff by auto with not_prime_1 p pa show False by (auto simp add: gcd_nat.order_iff) qed with y show ?thesis by blast qed lemma cong_unique_inverse_prime: fixes p :: nat assumes "prime p" and "0 < x" and "x < p" shows "\!y. 0 < y \ y < p \ [x * y = 1] (mod p)" by (rule cong_solve_unique_nontrivial) (use assms in simp_all) lemma chinese_remainder_coprime_unique: fixes a :: nat assumes ab: "coprime a b" and az: "a \ 0" and bz: "b \ 0" and ma: "coprime m a" and nb: "coprime n b" shows "\!x. coprime x (a * b) \ x < a * b \ [x = m] (mod a) \ [x = n] (mod b)" proof - let ?P = "\x. x < a * b \ [x = m] (mod a) \ [x = n] (mod b)" from binary_chinese_remainder_unique_nat[OF ab az bz] obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" "\y. ?P y \ y = x" by blast from ma nb x have "coprime x a" "coprime x b" using cong_imp_coprime cong_sym by blast+ then have "coprime x (a*b)" by simp with x show ?thesis by blast qed subsection \Lucas's theorem\ lemma lucas_coprime_lemma: fixes n :: nat assumes m: "m \ 0" and am: "[a^m = 1] (mod n)" shows "coprime a n" proof - consider "n = 1" | "n = 0" | "n > 1" by arith then show ?thesis proof cases case 1 then show ?thesis by simp next case 2 with am m show ?thesis by simp next case 3 from m obtain m' where m': "m = Suc m'" by (cases m) blast+ have "d = 1" if d: "d dvd a" "d dvd n" for d proof - from am mod_less[OF \n > 1\] have am1: "a^m mod n = 1" by (simp add: cong_def) from dvd_mult2[OF d(1), of "a^m'"] have dam: "d dvd a^m" by (simp add: m') from dvd_mod_iff[OF d(2), of "a^m"] dam am1 show ?thesis by simp qed then show ?thesis by (auto intro: coprimeI) qed qed lemma lucas_weak: fixes n :: nat assumes n: "n \ 2" and an: "[a ^ (n - 1) = 1] (mod n)" and nm: "\m. 0 < m \ m < n - 1 \ \ [a ^ m = 1] (mod n)" shows "prime n" proof (rule totient_imp_prime) show "totient n = n - 1" proof (rule ccontr) have "[a ^ totient n = 1] (mod n)" by (rule euler_theorem, rule lucas_coprime_lemma [of "n - 1"]) (use n an in auto) moreover assume "totient n \ n - 1" then have "totient n > 0" "totient n < n - 1" using \n \ 2\ and totient_less[of n] by simp_all ultimately show False using nm by auto qed qed (use n in auto) lemma nat_exists_least_iff: "(\(n::nat). P n) \ (\n. P n \ (\m < n. \ P m))" by (metis ex_least_nat_le not_less0) lemma nat_exists_least_iff': "(\(n::nat). P n) \ P (Least P) \ (\m < (Least P). \ P m)" (is "?lhs \ ?rhs") proof show ?lhs if ?rhs using that by blast show ?rhs if ?lhs proof - from \?lhs\ obtain n where n: "P n" by blast let ?x = "Least P" have "\ P m" if "m < ?x" for m by (rule not_less_Least[OF that]) with LeastI_ex[OF \?lhs\] show ?thesis by blast qed qed theorem lucas: assumes n2: "n \ 2" and an1: "[a^(n - 1) = 1] (mod n)" and pn: "\p. prime p \ p dvd n - 1 \ [a^((n - 1) div p) \ 1] (mod n)" shows "prime n" proof- from n2 have n01: "n \ 0" "n \ 1" "n - 1 \ 0" by arith+ from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime an1 have an: "coprime a n" "coprime (a ^ (n - 1)) n" using \n \ 2\ by simp_all have False if H0: "\m. 0 < m \ m < n - 1 \ [a ^ m = 1] (mod n)" (is "\m. ?P m") proof - from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\k ?P k" by blast have False if nm1: "(n - 1) mod m > 0" proof - from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast let ?y = "a^ ((n - 1) div m * m)" note mdeq = div_mult_mod_eq[of "(n - 1)" m] have yn: "coprime ?y n" using an(1) by (cases "(n - Suc 0) div m * m = 0") auto have "?y mod n = (a^m)^((n - 1) div m) mod n" by (simp add: algebra_simps power_mult) also have "\ = (a^m mod n)^((n - 1) div m) mod n" using power_mod[of "a^m" n "(n - 1) div m"] by simp also have "\ = 1" using m(3)[unfolded cong_def onen] onen by (metis power_one) finally have *: "?y mod n = 1" . have **: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)" using an1[unfolded cong_def onen] onen div_mult_mod_eq[of "(n - 1)" m, symmetric] by (simp add:power_add[symmetric] cong_def * del: One_nat_def) have "[a ^ ((n - 1) mod m) = 1] (mod n)" by (metis cong_mult_rcancel_nat mult.commute ** yn) with m(4)[rule_format, OF th0] nm1 less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] show ?thesis by blast qed then have "(n - 1) mod m = 0" by auto then have mn: "m dvd n - 1" by presburger then obtain r where r: "n - 1 = m * r" unfolding dvd_def by blast from n01 r m(2) have r01: "r \ 0" "r \ 1" by auto obtain p where p: "prime p" "p dvd r" by (metis prime_factor_nat r01(2)) then have th: "prime p \ p dvd n - 1" unfolding r by (auto intro: dvd_mult) from r have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" by (simp add: power_mult) also have "\ = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp also have "\ = ((a^m)^(r div p)) mod n" by (simp add: power_mult) also have "\ = ((a^m mod n)^(r div p)) mod n" using power_mod .. also from m(3) onen have "\ = 1" by (simp add: cong_def) finally have "[(a ^ ((n - 1) div p))= 1] (mod n)" using onen by (simp add: cong_def) with pn th show ?thesis by blast qed then have "\m. 0 < m \ m < n - 1 \ \ [a ^ m = 1] (mod n)" by blast then show ?thesis by (rule lucas_weak[OF n2 an1]) qed subsection \Definition of the order of a number mod \n\\ definition "ord n a = (if coprime n a then Least (\d. d > 0 \ [a ^d = 1] (mod n)) else 0)" text \This has the expected properties.\ lemma coprime_ord: fixes n::nat assumes "coprime n a" shows "ord n a > 0 \ [a ^(ord n a) = 1] (mod n) \ (\m. 0 < m \ m < ord n a \ [a^ m \ 1] (mod n))" proof- let ?P = "\d. 0 < d \ [a ^ d = 1] (mod n)" from bigger_prime[of a] obtain p where p: "prime p" "a < p" by blast from assms have o: "ord n a = Least ?P" by (simp add: ord_def) have ex: "\m>0. ?P m" proof (cases "n \ 2") case True moreover from assms have "coprime a n" by (simp add: ac_simps) then have "[a ^ totient n = 1] (mod n)" by (rule euler_theorem) ultimately show ?thesis by (auto intro: exI [where x = "totient n"]) next case False then have "n = 0 \ n = 1" by auto with assms show ?thesis by auto qed from nat_exists_least_iff'[of ?P] ex assms show ?thesis unfolding o[symmetric] by auto qed text \With the special value \0\ for non-coprime case, it's more convenient.\ lemma ord_works: "[a ^ (ord n a) = 1] (mod n) \ (\m. 0 < m \ m < ord n a \ \ [a^ m = 1] (mod n))" for n :: nat by (cases "coprime n a") (use coprime_ord[of n a] in \auto simp add: ord_def cong_def\) lemma ord: "[a^(ord n a) = 1] (mod n)" for n :: nat using ord_works by blast lemma ord_minimal: "0 < m \ m < ord n a \ \ [a^m = 1] (mod n)" for n :: nat using ord_works by blast lemma ord_eq_0: "ord n a = 0 \ \ coprime n a" for n :: nat by (cases "coprime n a") (simp add: coprime_ord, simp add: ord_def) lemma divides_rexp: "x dvd y \ x dvd (y ^ Suc n)" for x y :: nat by (simp add: dvd_mult2[of x y]) lemma ord_divides:"[a ^ d = 1] (mod n) \ ord n a dvd d" (is "?lhs \ ?rhs") for n :: nat proof assume ?rhs then obtain k where "d = ord n a * k" unfolding dvd_def by blast then have "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)" by (simp add : cong_def power_mult power_mod) also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)" using ord[of a n, unfolded cong_def] by (simp add: cong_def power_mod) finally show ?lhs . next assume ?lhs show ?rhs proof (cases "coprime n a") case prem: False then have o: "ord n a = 0" by (simp add: ord_def) show ?thesis proof (cases d) case 0 with o prem show ?thesis by (simp add: cong_def) next case (Suc d') then have d0: "d \ 0" by simp from prem obtain p where p: "p dvd n" "p dvd a" "p \ 1" by (auto elim: not_coprimeE) from \?lhs\ obtain q1 q2 where q12: "a ^ d + n * q1 = 1 + n * q2" using prem d0 lucas_coprime_lemma by (auto elim: not_coprimeE simp add: ac_simps) then have "a ^ d + n * q1 - n * q2 = 1" by simp with dvd_diff_nat [OF dvd_add [OF divides_rexp]] dvd_mult2 Suc p have "p dvd 1" by metis with p(3) have False by simp then show ?thesis .. qed next case H: True let ?o = "ord n a" let ?q = "d div ord n a" let ?r = "d mod ord n a" have eqo: "[(a^?o)^?q = 1] (mod n)" using cong_pow ord_works by fastforce from H have onz: "?o \ 0" by (simp add: ord_eq_0) then have opos: "?o > 0" by simp from div_mult_mod_eq[of d "ord n a"] \?lhs\ have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: cong_def mult.commute) then have "[(a^?o)^?q * (a^?r) = 1] (mod n)" by (simp add: cong_def power_mult[symmetric] power_add[symmetric]) then have th: "[a^?r = 1] (mod n)" using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n] by (simp add: cong_def del: One_nat_def) (metis mod_mult_left_eq nat_mult_1) show ?thesis proof (cases "?r = 0") case True then show ?thesis by (simp add: dvd_eq_mod_eq_0) next case False with mod_less_divisor[OF opos, of d] have r0o:"?r >0 \ ?r < ?o" by simp from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th show ?thesis by blast qed qed qed lemma order_divides_totient: "ord n a dvd totient n" if "coprime n a" using that euler_theorem [of a n] by (simp add: ord_divides [symmetric] ac_simps) lemma order_divides_expdiff: fixes n::nat and a::nat assumes na: "coprime n a" shows "[a^d = a^e] (mod n) \ [d = e] (mod (ord n a))" proof - have th: "[a^d = a^e] (mod n) \ [d = e] (mod (ord n a))" if na: "coprime n a" and ed: "(e::nat) \ d" for n a d e :: nat proof - from na ed have "\c. d = e + c" by presburger then obtain c where c: "d = e + c" .. from na have an: "coprime a n" by (simp add: ac_simps) then have aen: "coprime (a ^ e) n" by (cases "e > 0") simp_all from an have acn: "coprime (a ^ c) n" by (cases "c > 0") simp_all from c have "[a^d = a^e] (mod n) \ [a^(e + c) = a^(e + 0)] (mod n)" by simp also have "\ \ [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add) also have "\ \ [a ^ c = 1] (mod n)" using cong_mult_lcancel_nat [OF aen, of "a^c" "a^0"] by simp also have "\ \ ord n a dvd c" by (simp only: ord_divides) also have "\ \ [e + c = e + 0] (mod ord n a)" by (auto simp add: cong_altdef_nat) finally show ?thesis using c by simp qed consider "e \ d" | "d \ e" by arith then show ?thesis proof cases case 1 with na show ?thesis by (rule th) next case 2 from th[OF na this] show ?thesis by (metis cong_sym) qed qed lemma ord_not_coprime [simp]: "\coprime n a \ ord n a = 0" by (simp add: ord_def) lemma ord_1 [simp]: "ord 1 n = 1" proof - have "(LEAST k. k > 0) = (1 :: nat)" by (rule Least_equality) auto thus ?thesis by (simp add: ord_def) qed lemma ord_1_right [simp]: "ord (n::nat) 1 = 1" using ord_divides[of 1 1 n] by simp lemma ord_Suc_0_right [simp]: "ord (n::nat) (Suc 0) = 1" using ord_divides[of 1 1 n] by simp lemma ord_0_nat [simp]: "ord 0 (n :: nat) = (if n = 1 then 1 else 0)" proof - have "(LEAST k. k > 0) = (1 :: nat)" by (rule Least_equality) auto thus ?thesis by (auto simp: ord_def) qed lemma ord_0_right_nat [simp]: "ord (n :: nat) 0 = (if n = 1 then 1 else 0)" proof - have "(LEAST k. k > 0) = (1 :: nat)" by (rule Least_equality) auto thus ?thesis by (auto simp: ord_def) qed lemma ord_divides': "[a ^ d = Suc 0] (mod n) = (ord n a dvd d)" using ord_divides[of a d n] by simp lemma ord_Suc_0 [simp]: "ord (Suc 0) n = 1" using ord_1[where 'a = nat] by (simp del: ord_1) lemma ord_mod [simp]: "ord n (k mod n) = ord n k" by (cases "n = 0") (auto simp add: ord_def cong_def power_mod) lemma ord_gt_0_iff [simp]: "ord (n::nat) x > 0 \ coprime n x" using ord_eq_0[of n x] by auto lemma ord_eq_Suc_0_iff: "ord n (x::nat) = Suc 0 \ [x = 1] (mod n)" using ord_divides[of x 1 n] by (auto simp: ord_divides') lemma ord_cong: assumes "[k1 = k2] (mod n)" shows "ord n k1 = ord n k2" proof - have "ord n (k1 mod n) = ord n (k2 mod n)" by (simp only: assms[unfolded cong_def]) thus ?thesis by simp qed lemma ord_nat_code [code_unfold]: "ord n a = (if n = 0 then if a = 1 then 1 else 0 else if coprime n a then Min (Set.filter (\k. [a ^ k = 1] (mod n)) {0<..n}) else 0)" proof (cases "coprime n a \ n > 0") case True define A where "A = {k\{0<..n}. [a ^ k = 1] (mod n)}" define k where "k = (LEAST k. k > 0 \ [a ^ k = 1] (mod n))" have totient: "totient n \ A" using euler_theorem[of a n] True by (auto simp: A_def coprime_commute intro!: Nat.gr0I totient_le) moreover have "finite A" by (auto simp: A_def) ultimately have *: "Min A \ A" and "\y. y \ A \ Min A \ y" by (auto intro: Min_in) have "k > 0 \ [a ^ k = 1] (mod n)" unfolding k_def by (rule LeastI[of _ "totient n"]) (use totient in \auto simp: A_def\) moreover have "k \ totient n" unfolding k_def by (intro Least_le) (use totient in \auto simp: A_def\) ultimately have "k \ A" using totient_le[of n] by (auto simp: A_def) hence "Min A \ k" by (intro Min_le) (auto simp: \finite A\) moreover from * have "k \ Min A" unfolding k_def by (intro Least_le) (auto simp: A_def) ultimately show ?thesis using True by (simp add: ord_def k_def A_def Set.filter_def) qed auto theorem ord_modulus_mult_coprime: fixes x :: nat assumes "coprime m n" shows "ord (m * n) x = lcm (ord m x) (ord n x)" proof (intro dvd_antisym) have "[x ^ lcm (ord m x) (ord n x) = 1] (mod (m * n))" using assms by (intro coprime_cong_mult_nat assms) (auto simp: ord_divides') thus "ord (m * n) x dvd lcm (ord m x) (ord n x)" by (simp add: ord_divides') next show "lcm (ord m x) (ord n x) dvd ord (m * n) x" proof (intro lcm_least) show "ord m x dvd ord (m * n) x" using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 m n] assms by (simp add: ord_divides') show "ord n x dvd ord (m * n) x" using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 n m] assms by (simp add: ord_divides' mult.commute) qed qed corollary ord_modulus_prod_coprime: assumes "finite A" "\i j. i \ A \ j \ A \ i \ j \ coprime (f i) (f j)" shows "ord (\i\A. f i :: nat) x = (LCM i\A. ord (f i) x)" using assms by (induction A rule: finite_induct) (simp, simp, subst ord_modulus_mult_coprime, auto intro!: prod_coprime_right) lemma ord_power_aux: fixes m x k a :: nat defines "l \ ord m a" shows "ord m (a ^ k) * gcd k l = l" proof (rule dvd_antisym) have "[a ^ lcm k l = 1] (mod m)" unfolding ord_divides by (simp add: l_def) also have "lcm k l = k * (l div gcd k l)" by (simp add: lcm_nat_def div_mult_swap) finally have "ord m (a ^ k) dvd l div gcd k l" unfolding ord_divides [symmetric] by (simp add: power_mult [symmetric]) thus "ord m (a ^ k) * gcd k l dvd l" by (cases "l = 0") (auto simp: dvd_div_iff_mult) have "[(a ^ k) ^ ord m (a ^ k) = 1] (mod m)" by (rule ord) also have "(a ^ k) ^ ord m (a ^ k) = a ^ (k * ord m (a ^ k))" by (simp add: power_mult) finally have "ord m a dvd k * ord m (a ^ k)" by (simp add: ord_divides') hence "l dvd gcd (k * ord m (a ^ k)) (l * ord m (a ^ k))" by (intro gcd_greatest dvd_triv_left) (auto simp: l_def ord_divides') also have "gcd (k * ord m (a ^ k)) (l * ord m (a ^ k)) = ord m (a ^ k) * gcd k l" by (subst gcd_mult_distrib_nat) (auto simp: mult_ac) finally show "l dvd ord m (a ^ k) * gcd k l" . qed theorem ord_power: "coprime m a \ ord m (a ^ k :: nat) = ord m a div gcd k (ord m a)" using ord_power_aux[of m a k] by (metis div_mult_self_is_m gcd_pos_nat ord_eq_0) lemma inj_power_mod: assumes "coprime n (a :: nat)" shows "inj_on (\k. a ^ k mod n) {.. {.. {.. = a ^ k * a ^ (l - k)" by (simp add: power_add) also have "[\ = a ^ l * a ^ (l - k)] (mod n)" using that by (intro cong_mult) auto finally have "[a ^ l * a ^ (l - k) = a ^ l * 1] (mod n)" by (simp add: cong_sym_eq) with assms have "[a ^ (l - k) = 1] (mod n)" by (subst (asm) cong_mult_lcancel_nat) (auto simp: coprime_commute) hence "ord n a dvd l - k" by (simp add: ord_divides') from dvd_imp_le[OF this] and \l < ord n a\ have "l - k = 0" by (cases "l - k = 0") auto with \k < l\ show "k = l" by simp qed from this[of k l] and this[of l k] and * show "k = l" by (cases k l rule: linorder_cases) (auto simp: cong_def) qed lemma ord_eq_2_iff: "ord n (x :: nat) = 2 \ [x \ 1] (mod n) \ [x\<^sup>2 = 1] (mod n)" proof assume x: "[x \ 1] (mod n) \ [x\<^sup>2 = 1] (mod n)" hence "coprime n x" by (metis coprime_commute lucas_coprime_lemma zero_neq_numeral) with x have "ord n x dvd 2" "ord n x \ 1" "ord n x > 0" by (auto simp: ord_divides' ord_eq_Suc_0_iff) thus "ord n x = 2" by (auto dest!: dvd_imp_le simp del: ord_gt_0_iff) qed (use ord_divides[of _ 2] ord_divides[of _ 1] in auto) lemma square_mod_8_eq_1_iff: "[x\<^sup>2 = 1] (mod 8) \ odd (x :: nat)" proof - have "[x\<^sup>2 = 1] (mod 8) \ ((x mod 8)\<^sup>2 mod 8 = 1)" by (simp add: power_mod cong_def) also have "\ \ x mod 8 \ {1, 3, 5, 7}" proof assume x: "(x mod 8)\<^sup>2 mod 8 = 1" have "x mod 8 \ {..<8}" by simp also have "{..<8} = {0, 1, 2, 3, 4, 5, 6, 7::nat}" by (simp add: lessThan_nat_numeral lessThan_Suc insert_commute) finally have x_cases: "x mod 8 \ {0, 1, 2, 3, 4, 5, 6, 7}" . from x have "x mod 8 \ {0, 2, 4, 6}" using x by (auto intro: Nat.gr0I) with x_cases show "x mod 8 \ {1, 3, 5, 7}" by simp qed auto also have "\ \ odd (x mod 8)" by (auto elim!: oddE) also have "\ \ odd x" by presburger finally show ?thesis . qed lemma ord_twopow_aux: assumes "k \ 3" and "odd (x :: nat)" shows "[x ^ (2 ^ (k - 2)) = 1] (mod (2 ^ k))" using assms(1) proof (induction k rule: dec_induct) case base from assms have "[x\<^sup>2 = 1] (mod 8)" by (subst square_mod_8_eq_1_iff) auto thus ?case by simp next case (step k) define k' where "k' = k - 2" have k: "k = Suc (Suc k')" using \k \ 3\ by (simp add: k'_def) from \k \ 3\ have "2 * k \ Suc k" by presburger from \odd x\ have "x > 0" by (intro Nat.gr0I) auto from step.IH have "2 ^ k dvd (x ^ (2 ^ (k - 2)) - 1)" by (rule cong_to_1_nat) then obtain t where "x ^ (2 ^ (k - 2)) - 1 = t * 2 ^ k" by auto hence "x ^ (2 ^ (k - 2)) = t * 2 ^ k + 1" by (metis \0 < x\ add.commute add_diff_inverse_nat less_one neq0_conv power_eq_0_iff) hence "(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2" by (rule arg_cong) hence "[(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2] (mod (2 ^ Suc k))" by simp also have "(x ^ (2 ^ (k - 2))) ^ 2 = x ^ (2 ^ (k - 1))" by (simp_all add: power_even_eq[symmetric] power_mult k ) also have "(t * 2 ^ k + 1) ^ 2 = t\<^sup>2 * 2 ^ (2 * k) + t * 2 ^ Suc k + 1" by (subst power2_eq_square) (auto simp: algebra_simps k power2_eq_square[of t] power_even_eq[symmetric] power_add [symmetric]) also have "[\ = 0 + 0 + 1] (mod 2 ^ Suc k)" using \2 * k \ Suc k\ by (intro cong_add) (auto simp: cong_0_iff intro: dvd_mult[OF le_imp_power_dvd] simp del: power_Suc) finally show ?case by simp qed lemma ord_twopow_3_5: assumes "k \ 3" "x mod 8 \ {3, 5 :: nat}" shows "ord (2 ^ k) x = 2 ^ (k - 2)" using assms(1) proof (induction k rule: less_induct) have "x mod 8 = 3 \ x mod 8 = 5" using assms by auto hence "odd x" by presburger case (less k) from \k \ 3\ consider "k = 3" | "k = 4" | "k \ 5" by force thus ?case proof cases case 1 thus ?thesis using assms by (auto simp: ord_eq_2_iff cong_def simp flip: power_mod[of x]) next case 2 from assms have "x mod 8 = 3 \ x mod 8 = 5" by auto - hence x': "x mod 16 = 3 \ x mod 16 = 5 \ x mod 16 = 11 \ x mod 16 = 13" - using mod_double_modulus[of 8 x] by auto + then have x': "x mod 16 = 3 \ x mod 16 = 5 \ x mod 16 = 11 \ x mod 16 = 13" + using mod_double_modulus [of 8 x] by auto hence "[x ^ 4 = 1] (mod 16)" using assms by (auto simp: cong_def simp flip: power_mod[of x]) hence "ord 16 x dvd 2\<^sup>2" by (simp add: ord_divides') then obtain l where l: "ord 16 x = 2 ^ l" "l \ 2" by (subst (asm) divides_primepow_nat) auto have "[x ^ 2 \ 1] (mod 16)" using x' by (auto simp: cong_def simp flip: power_mod[of x]) hence "\ord 16 x dvd 2" by (simp add: ord_divides') with l have "l = 2" using le_imp_power_dvd[of l 1 2] by (cases "l \ 1") auto with l show ?thesis by (simp add: \k = 4\) next case 3 define k' where "k' = k - 2" have k': "k' \ 2" and [simp]: "k = Suc (Suc k')" using 3 by (simp_all add: k'_def) have IH: "ord (2 ^ k') x = 2 ^ (k' - 2)" "ord (2 ^ Suc k') x = 2 ^ (k' - 1)" using less.IH[of k'] less.IH[of "Suc k'"] 3 by simp_all from IH have cong: "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ k'))" by (simp_all add: ord_divides') have notcong: "[x ^ (2 ^ (k' - 2)) \ 1] (mod (2 ^ Suc k'))" proof assume "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ Suc k'))" hence "ord (2 ^ Suc k') x dvd 2 ^ (k' - 2)" by (simp add: ord_divides') also have "ord (2 ^ Suc k') x = 2 ^ (k' - 1)" using IH by simp finally have "k' - 1 \ k' - 2" by (rule power_dvd_imp_le) auto with \k' \ 2\ show False by simp qed have "2 ^ k' + 1 < 2 ^ k' + (2 ^ k' :: nat)" using one_less_power[of "2::nat" k'] k' by (intro add_strict_left_mono) auto with cong notcong have cong': "x ^ (2 ^ (k' - 2)) mod 2 ^ Suc k' = 1 + 2 ^ k'" - using mod_double_modulus[of "2 ^ k'" "x ^ 2 ^ (k' - 2)"] k' by (auto simp: cong_def) + using mod_double_modulus [of "2 ^ k'" "x ^ 2 ^ (k' - 2)"] k' by (auto simp: cong_def) hence "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' \ x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'" - using mod_double_modulus[of "2 ^ Suc k'" "x ^ 2 ^ (k' - 2)"] by auto + using mod_double_modulus [of "2 ^ Suc k'" "x ^ 2 ^ (k' - 2)"] by auto hence eq: "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" proof assume *: "x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'" have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)" by simp also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'] (mod 2 ^ k)" by (subst *) auto finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)" by (rule cong_pow) hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)" by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc) also have "Suc (k' - 2) = k' - 1" using k' by simp also have "(1 + 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ (2 * k')" by (subst power2_eq_square) (simp add: algebra_simps flip: power_add) also have "(2 ^ k :: nat) dvd 2 ^ (2 * k')" using k' by (intro le_imp_power_dvd) auto hence "[1 + 2 ^ (k - 1) + 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + (0 :: nat)] (mod 2 ^ k)" by (intro cong_add) (auto simp: cong_0_iff) finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" by simp next assume *: "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'" have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)" by simp also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 3 * 2 ^ k'] (mod 2 ^ k)" by (subst *) auto finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)" by (rule cong_pow) hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)" by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc) also have "Suc (k' - 2) = k' - 1" using k' by simp also have "(1 + 3 * 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k')" by (subst power2_eq_square) (simp add: algebra_simps flip: power_add) also have "(2 ^ k :: nat) dvd 9 * 2 ^ (2 * k')" using k' by (intro dvd_mult le_imp_power_dvd) auto hence "[1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + 0 + (0 :: nat)] (mod 2 ^ k)" by (intro cong_add) (auto simp: cong_0_iff) finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)" by simp qed have notcong': "[x ^ 2 ^ (k - 3) \ 1] (mod 2 ^ k)" proof assume "[x ^ 2 ^ (k - 3) = 1] (mod 2 ^ k)" hence "[x ^ 2 ^ (k' - 1) - x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1) - 1] (mod 2 ^ k)" by (intro cong_diff_nat eq) auto hence "[2 ^ (k - 1) = (0 :: nat)] (mod 2 ^ k)" by (simp add: cong_sym_eq) hence "2 ^ k dvd 2 ^ (k - 1)" by (simp add: cong_0_iff) hence "k \ k - 1" by (rule power_dvd_imp_le) auto thus False by simp qed have "[x ^ 2 ^ (k - 2) = 1] (mod 2 ^ k)" using ord_twopow_aux[of k x] \odd x\ \k \ 3\ by simp hence "ord (2 ^ k) x dvd 2 ^ (k - 2)" by (simp add: ord_divides') then obtain l where l: "l \ k - 2" "ord (2 ^ k) x = 2 ^ l" using divides_primepow_nat[of 2 "ord (2 ^ k) x" "k - 2"] by auto from notcong' have "\ord (2 ^ k) x dvd 2 ^ (k - 3)" by (simp add: ord_divides') with l have "l = k - 2" using le_imp_power_dvd[of l "k - 3" 2] by (cases "l \ k - 3") auto with l show ?thesis by simp qed qed lemma ord_4_3 [simp]: "ord 4 (3::nat) = 2" proof - have "[3 ^ 2 = (1 :: nat)] (mod 4)" by (simp add: cong_def) hence "ord 4 (3::nat) dvd 2" by (subst (asm) ord_divides) auto hence "ord 4 (3::nat) \ 2" by (intro dvd_imp_le) auto moreover have "ord 4 (3::nat) \ 1" by (auto simp: ord_eq_Suc_0_iff cong_def) moreover have "ord 4 (3::nat) \ 0" by (auto simp: gcd_non_0_nat coprime_iff_gcd_eq_1) ultimately show "ord 4 (3 :: nat) = 2" by linarith qed lemma elements_with_ord_1: "n > 0 \ {x\totatives n. ord n x = Suc 0} = {1}" by (auto simp: ord_eq_Suc_0_iff cong_def totatives_less) lemma residue_prime_has_primroot: fixes p :: nat assumes "prime p" shows "\a\totatives p. ord p a = p - 1" proof - from residue_prime_mult_group_has_gen[OF assms] obtain a where a: "a \ {1..p-1}" "{1..p-1} = {a ^ i mod p |i. i \ UNIV}" by blast from a have "coprime p a" using a assms by (intro prime_imp_coprime) (auto dest: dvd_imp_le) with a(1) have "a \ totatives p" by (auto simp: totatives_def coprime_commute) have "p - 1 = card {1..p-1}" by simp also have "{1..p-1} = {a ^ i mod p |i. i \ UNIV}" by fact also have "{a ^ i mod p |i. i \ UNIV} = (\i. a ^ i mod p) ` {.. {a ^ i mod p |i. i \ UNIV}" then obtain i where [simp]: "x = a ^ i mod p" by auto have "[a ^ i = a ^ (i mod ord p a)] (mod p)" using \coprime p a\ by (subst order_divides_expdiff) auto hence "\j. a ^ i mod p = a ^ j mod p \ j < ord p a" using \coprime p a\ by (intro exI[of _ "i mod ord p a"]) (auto simp: cong_def) thus "x \ (\i. a ^ i mod p) ` {.. = ord p a" using inj_power_mod[OF \coprime p a\] by (subst card_image) auto finally show ?thesis using \a \ totatives p\ by auto qed subsection \Another trivial primality characterization\ lemma prime_prime_factor: "prime n \ n \ 1 \ (\p. prime p \ p dvd n \ p = n)" (is "?lhs \ ?rhs") for n :: nat proof (cases "n = 0 \ n = 1") case True then show ?thesis by (metis bigger_prime dvd_0_right not_prime_1 not_prime_0) next case False show ?thesis proof assume "prime n" then show ?rhs by (metis not_prime_1 prime_nat_iff) next assume ?rhs with False show "prime n" by (auto simp: prime_nat_iff) (metis One_nat_def prime_factor_nat prime_nat_iff) qed qed lemma prime_divisor_sqrt: "prime n \ n \ 1 \ (\d. d dvd n \ d\<^sup>2 \ n \ d = 1)" for n :: nat proof - consider "n = 0" | "n = 1" | "n \ 0" "n \ 1" by blast then show ?thesis proof cases case 1 then show ?thesis by simp next case 2 then show ?thesis by simp next case n: 3 then have np: "n > 1" by arith { fix d assume d: "d dvd n" "d\<^sup>2 \ n" and H: "\m. m dvd n \ m = 1 \ m = n" from H d have d1n: "d = 1 \ d = n" by blast then have "d = 1" proof assume dn: "d = n" from n have "n\<^sup>2 > n * 1" by (simp add: power2_eq_square) with dn d(2) show ?thesis by simp qed } moreover { fix d assume d: "d dvd n" and H: "\d'. d' dvd n \ d'\<^sup>2 \ n \ d' = 1" from d n have "d \ 0" by (metis dvd_0_left_iff) then have dp: "d > 0" by simp from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast from n dp e have ep:"e > 0" by simp from dp ep have "d\<^sup>2 \ n \ e\<^sup>2 \ n" by (auto simp add: e power2_eq_square mult_le_cancel_left) then have "d = 1 \ d = n" proof assume "d\<^sup>2 \ n" with H[rule_format, of d] d have "d = 1" by blast then show ?thesis .. next assume h: "e\<^sup>2 \ n" from e have "e dvd n" by (simp add: dvd_def mult.commute) with H[rule_format, of e] h have "e = 1" by simp with e have "d = n" by simp then show ?thesis .. qed } ultimately show ?thesis unfolding prime_nat_iff using np n(2) by blast qed qed lemma prime_prime_factor_sqrt: "prime (n::nat) \ n \ 0 \ n \ 1 \ (\p. prime p \ p dvd n \ p\<^sup>2 \ n)" (is "?lhs \?rhs") proof - consider "n = 0" | "n = 1" | "n \ 0" "n \ 1" by blast then show ?thesis proof cases case 1 then show ?thesis by (metis not_prime_0) next case 2 then show ?thesis by (metis not_prime_1) next case n: 3 show ?thesis proof assume ?lhs from this[unfolded prime_divisor_sqrt] n show ?rhs by (metis prime_prime_factor) next assume ?rhs { fix d assume d: "d dvd n" "d\<^sup>2 \ n" "d \ 1" then obtain p where p: "prime p" "p dvd d" by (metis prime_factor_nat) from d(1) n have dp: "d > 0" by (metis dvd_0_left neq0_conv) from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2) have "p\<^sup>2 \ n" unfolding power2_eq_square by arith with \?rhs\ n p(1) dvd_trans[OF p(2) d(1)] have False by blast } with n prime_divisor_sqrt show ?lhs by auto qed qed qed subsection \Pocklington theorem\ lemma pocklington_lemma: fixes p :: nat assumes n: "n \ 2" and nqr: "n - 1 = q * r" and an: "[a^ (n - 1) = 1] (mod n)" and aq: "\p. prime p \ p dvd q \ coprime (a ^ ((n - 1) div p) - 1) n" and pp: "prime p" and pn: "p dvd n" shows "[p = 1] (mod q)" proof - have p01: "p \ 0" "p \ 1" using pp by (auto intro: prime_gt_0_nat) obtain k where k: "a ^ (q * r) - 1 = n * k" by (metis an cong_to_1_nat dvd_def nqr) from pn[unfolded dvd_def] obtain l where l: "n = p * l" by blast have a0: "a \ 0" proof assume "a = 0" with n have "a^ (n - 1) = 0" by (simp add: power_0_left) with n an mod_less[of 1 n] show False by (simp add: power_0_left cong_def) qed with n nqr have aqr0: "a ^ (q * r) \ 0" by simp then have "(a ^ (q * r) - 1) + 1 = a ^ (q * r)" by simp with k l have "a ^ (q * r) = p * l * k + 1" by simp then have "a ^ (r * q) + p * 0 = 1 + p * (l * k)" by (simp add: ac_simps) then have odq: "ord p (a^r) dvd q" unfolding ord_divides[symmetric] power_mult[symmetric] by (metis an cong_dvd_modulus_nat mult.commute nqr pn) from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast have d1: "d = 1" proof (rule ccontr) assume d1: "d \ 1" obtain P where P: "prime P" "P dvd d" by (metis d1 prime_factor_nat) from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast from P(1) have P0: "P \ 0" by (metis not_prime_0) from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast from d s t P0 have s': "ord p (a^r) * t = s" by (metis mult.commute mult_cancel1 mult.assoc) have "ord p (a^r) * t*r = r * ord p (a^r) * t" by (metis mult.assoc mult.commute) then have exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t" by (simp only: power_mult) then have "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)" by (metis cong_pow ord power_one) then have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by (metis cong_to_1_nat exps) from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp with caP have "coprime (a ^ (ord p (a ^ r) * t * r) - 1) n" by simp with p01 pn pd0 coprime_common_divisor [of _ n p] show False by auto qed with d have o: "ord p (a^r) = q" by simp from pp totient_prime [of p] have totient_eq: "totient p = p - 1" by simp { fix d assume d: "d dvd p" "d dvd a" "d \ 1" from pp[unfolded prime_nat_iff] d have dp: "d = p" by blast from n have "n \ 0" by simp then have False using d dp pn an by auto (metis One_nat_def Suc_lessI \1 < p \ (\m. m dvd p \ m = 1 \ m = p)\ \a ^ (q * r) = p * l * k + 1\ add_diff_cancel_left' dvd_diff_nat dvd_power dvd_triv_left gcd_nat.trans nat_dvd_not_less nqr zero_less_diff zero_less_one) } then have cpa: "coprime p a" by (auto intro: coprimeI) then have arp: "coprime (a ^ r) p" by (cases "r > 0") (simp_all add: ac_simps) from euler_theorem [OF arp, simplified ord_divides] o totient_eq have "q dvd (p - 1)" by simp then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast have "p \ 0" by (metis p01(1)) with d have "p + q * 0 = 1 + q * d" by simp then show ?thesis by (metis cong_iff_lin_nat mult.commute) qed theorem pocklington: assumes n: "n \ 2" and nqr: "n - 1 = q * r" and sqr: "n \ q\<^sup>2" and an: "[a^ (n - 1) = 1] (mod n)" and aq: "\p. prime p \ p dvd q \ coprime (a^ ((n - 1) div p) - 1) n" shows "prime n" unfolding prime_prime_factor_sqrt[of n] proof - let ?ths = "n \ 0 \ n \ 1 \ (\p. prime p \ p dvd n \ p\<^sup>2 \ n)" from n have n01: "n \ 0" "n \ 1" by arith+ { fix p assume p: "prime p" "p dvd n" "p\<^sup>2 \ n" from p(3) sqr have "p^(Suc 1) \ q^(Suc 1)" by (simp add: power2_eq_square) then have pq: "p \ q" by (metis le0 power_le_imp_le_base) from pocklington_lemma[OF n nqr an aq p(1,2)] have *: "q dvd p - 1" by (metis cong_to_1_nat) have "p - 1 \ 0" using prime_ge_2_nat [OF p(1)] by arith with pq * have False by (simp add: nat_dvd_not_less) } with n01 show ?ths by blast qed text \Variant for application, to separate the exponentiation.\ lemma pocklington_alt: assumes n: "n \ 2" and nqr: "n - 1 = q * r" and sqr: "n \ q\<^sup>2" and an: "[a^ (n - 1) = 1] (mod n)" and aq: "\p. prime p \ p dvd q \ (\b. [a^((n - 1) div p) = b] (mod n) \ coprime (b - 1) n)" shows "prime n" proof - { fix p assume p: "prime p" "p dvd q" from aq[rule_format] p obtain b where b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast have a0: "a \ 0" proof assume a0: "a = 0" from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto then show False using n by (simp add: cong_def dvd_eq_mod_eq_0[symmetric]) qed then have a1: "a \ 1" by arith from one_le_power[OF a1] have ath: "1 \ a ^ ((n - 1) div p)" . have b0: "b \ 0" proof assume b0: "b = 0" from p(2) nqr have "(n - 1) mod p = 0" by (metis mod_0 mod_mod_cancel mod_mult_self1_is_0) with div_mult_mod_eq[of "n - 1" p] have "(n - 1) div p * p= n - 1" by auto then have eq: "(a^((n - 1) div p))^p = a^(n - 1)" by (simp only: power_mult[symmetric]) have "p - 1 \ 0" using prime_ge_2_nat [OF p(1)] by arith then have pS: "Suc (p - 1) = p" by arith from b have d: "n dvd a^((n - 1) div p)" unfolding b0 by auto from divides_rexp[OF d, of "p - 1"] pS eq cong_dvd_iff [OF an] n show False by simp qed then have b1: "b \ 1" by arith from cong_imp_coprime[OF Cong.cong_diff_nat[OF cong_sym [OF b(1)] cong_refl [of 1] b1]] ath b1 b nqr have "coprime (a ^ ((n - 1) div p) - 1) n" by simp } then have "\p. prime p \ p dvd q \ coprime (a ^ ((n - 1) div p) - 1) n " by blast then show ?thesis by (rule pocklington[OF n nqr sqr an]) qed subsection \Prime factorizations\ (* FIXME some overlap with material in UniqueFactorization, class unique_factorization *) definition "primefact ps n \ foldr (*) ps 1 = n \ (\p\ set ps. prime p)" lemma primefact: fixes n :: nat assumes n: "n \ 0" shows "\ps. primefact ps n" proof - obtain xs where xs: "mset xs = prime_factorization n" using ex_mset [of "prime_factorization n"] by blast from assms have "n = prod_mset (prime_factorization n)" by (simp add: prod_mset_prime_factorization) also have "\ = prod_mset (mset xs)" by (simp add: xs) also have "\ = foldr (*) xs 1" by (induct xs) simp_all finally have "foldr (*) xs 1 = n" .. moreover from xs have "\p\#mset xs. prime p" by auto ultimately have "primefact xs n" by (auto simp: primefact_def) then show ?thesis .. qed lemma primefact_contains: fixes p :: nat assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n" shows "p \ set ps" using pf p pn proof (induct ps arbitrary: p n) case Nil then show ?case by (auto simp: primefact_def) next case (Cons q qs) from Cons.prems[unfolded primefact_def] have q: "prime q" "q * foldr (*) qs 1 = n" "\p \set qs. prime p" and p: "prime p" "p dvd q * foldr (*) qs 1" by simp_all consider "p dvd q" | "p dvd foldr (*) qs 1" by (metis p prime_dvd_mult_eq_nat) then show ?case proof cases case 1 with p(1) q(1) have "p = q" unfolding prime_nat_iff by auto then show ?thesis by simp next case prem: 2 from q(3) have pqs: "primefact qs (foldr (*) qs 1)" by (simp add: primefact_def) from Cons.hyps[OF pqs p(1) prem] show ?thesis by simp qed qed lemma primefact_variant: "primefact ps n \ foldr (*) ps 1 = n \ list_all prime ps" by (auto simp add: primefact_def list_all_iff) text \Variant of Lucas theorem.\ lemma lucas_primefact: assumes n: "n \ 2" and an: "[a^(n - 1) = 1] (mod n)" and psn: "foldr (*) ps 1 = n - 1" and psp: "list_all (\p. prime p \ \ [a^((n - 1) div p) = 1] (mod n)) ps" shows "prime n" proof - { fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)" from psn psp have psn1: "primefact ps (n - 1)" by (auto simp add: list_all_iff primefact_variant) from p(3) primefact_contains[OF psn1 p(1,2)] psp have False by (induct ps) auto } with lucas[OF n an] show ?thesis by blast qed text \Variant of Pocklington theorem.\ lemma pocklington_primefact: assumes n: "n \ 2" and qrn: "q*r = n - 1" and nq2: "n \ q\<^sup>2" and arnb: "(a^r) mod n = b" and psq: "foldr (*) ps 1 = q" and bqn: "(b^q) mod n = 1" and psp: "list_all (\p. prime p \ coprime ((b^(q div p)) mod n - 1) n) ps" shows "prime n" proof - from bqn psp qrn have bqn: "a ^ (n - 1) mod n = 1" and psp: "list_all (\p. prime p \ coprime (a^(r *(q div p)) mod n - 1) n) ps" unfolding arnb[symmetric] power_mod by (simp_all add: power_mult[symmetric] algebra_simps) from n have n0: "n > 0" by arith from div_mult_mod_eq[of "a^(n - 1)" n] mod_less_divisor[OF n0, of "a^(n - 1)"] have an1: "[a ^ (n - 1) = 1] (mod n)" by (metis bqn cong_def mod_mod_trivial) have "coprime (a ^ ((n - 1) div p) - 1) n" if p: "prime p" "p dvd q" for p proof - from psp psq have pfpsq: "primefact ps q" by (auto simp add: primefact_variant list_all_iff) from psp primefact_contains[OF pfpsq p] have p': "coprime (a ^ (r * (q div p)) mod n - 1) n" by (simp add: list_all_iff) from p prime_nat_iff have p01: "p \ 0" "p \ 1" "p = Suc (p - 1)" by auto from div_mult1_eq[of r q p] p(2) have eq1: "r* (q div p) = (n - 1) div p" unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute) have ath: "a \ b \ a \ 0 \ 1 \ a \ 1 \ b" for a b :: nat by arith { assume "a ^ ((n - 1) div p) mod n = 0" then obtain s where s: "a ^ ((n - 1) div p) = n * s" by blast then have eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp from qrn[symmetric] have qn1: "q dvd n - 1" by (auto simp: dvd_def) from dvd_trans[OF p(2) qn1] have npp: "(n - 1) div p * p = n - 1" by simp with eq0 have "a ^ (n - 1) = (n * s) ^ p" by (simp add: power_mult[symmetric]) with bqn p01 have "1 = (n * s)^(Suc (p - 1)) mod n" by simp also have "\ = 0" by (simp add: mult.assoc) finally have False by simp } then have *: "a ^ ((n - 1) div p) mod n \ 0" by auto have "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)" by (simp add: cong_def) with ath[OF mod_less_eq_dividend *] have "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)" by (simp add: cong_diff_nat) then show ?thesis by (metis cong_imp_coprime eq1 p') qed with pocklington[OF n qrn[symmetric] nq2 an1] show ?thesis by blast qed end diff --git a/src/HOL/UNITY/Simple/Token.thy b/src/HOL/UNITY/Simple/Token.thy --- a/src/HOL/UNITY/Simple/Token.thy +++ b/src/HOL/UNITY/Simple/Token.thy @@ -1,134 +1,135 @@ (* Title: HOL/UNITY/Simple/Token.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge *) section \The Token Ring\ theory Token imports "../WFair" begin text\From Misra, "A Logic for Concurrent Programming" (1994), sections 5.2 and 13.2.\ subsection\Definitions\ datatype pstate = Hungry | Eating | Thinking \ \process states\ record state = token :: "nat" proc :: "nat => pstate" definition HasTok :: "nat => state set" where "HasTok i == {s. token s = i}" definition H :: "nat => state set" where "H i == {s. proc s i = Hungry}" definition E :: "nat => state set" where "E i == {s. proc s i = Eating}" definition T :: "nat => state set" where "T i == {s. proc s i = Thinking}" locale Token = fixes N and F and nodeOrder and "next" defines nodeOrder_def: "nodeOrder j == measure(%i. ((j+N)-i) mod N) \ {.. {.. (T i) co (T i \ H i)" and TR3: "F \ (H i) co (H i \ E i)" and TR4: "F \ (H i - HasTok i) co (H i)" and TR5: "F \ (HasTok i) co (HasTok i \ -(E i))" and TR6: "F \ (H i \ HasTok i) leadsTo (E i)" and TR7: "F \ (HasTok i) leadsTo (HasTok (next i))" lemma HasToK_partition: "[| s \ HasTok i; s \ HasTok j |] ==> i=j" by (unfold HasTok_def, auto) lemma not_E_eq: "(s \ E i) = (s \ H i | s \ T i)" apply (simp add: H_def E_def T_def) apply (cases "proc s i", auto) done context Token begin lemma token_stable: "F \ stable (-(E i) \ (HasTok i))" apply (unfold stable_def) apply (rule constrains_weaken) apply (rule constrains_Un [OF constrains_Un [OF TR2 TR4] TR5]) apply (auto simp add: not_E_eq) apply (simp_all add: H_def E_def T_def) done subsection\Progress under Weak Fairness\ lemma wf_nodeOrder: "wf(nodeOrder j)" apply (unfold nodeOrder_def) apply (rule wf_measure [THEN wf_subset], blast) done lemma nodeOrder_eq: "[| i ((next i, i) \ nodeOrder j) = (i \ j)" -apply (unfold nodeOrder_def next_def) -apply (auto simp add: mod_Suc mod_geq) -apply (auto split: nat_diff_split simp add: linorder_neq_iff mod_geq) -done + apply (cases \i < j\) + apply (auto simp add: nodeOrder_def next_def mod_Suc add.commute [of _ N]) + apply (simp only: diff_add_assoc mod_add_self1) + apply simp + done text\From "A Logic for Concurrent Programming", but not used in Chapter 4. Note the use of \cases\. Reasoning about leadsTo takes practice!\ lemma TR7_nodeOrder: "[| i F \ (HasTok i) leadsTo ({s. (token s, i) \ nodeOrder j} \ HasTok j)" apply (cases "i=j") apply (blast intro: subset_imp_leadsTo) apply (rule TR7 [THEN leadsTo_weaken_R]) apply (auto simp add: HasTok_def nodeOrder_eq) done text\Chapter 4 variant, the one actually used below.\ lemma TR7_aux: "[| ij |] ==> F \ (HasTok i) leadsTo {s. (token s, i) \ nodeOrder j}" apply (rule TR7 [THEN leadsTo_weaken_R]) apply (auto simp add: HasTok_def nodeOrder_eq) done lemma token_lemma: "({s. token s < N} \ token -` {m}) = (if mMisra's TR9: the token reaches an arbitrary node\ lemma leadsTo_j: "j F \ {s. token s < N} leadsTo (HasTok j)" apply (rule leadsTo_weaken_R) apply (rule_tac I = "-{j}" and f = token and B = "{}" in wf_nodeOrder [THEN bounded_induct]) apply (simp_all (no_asm_simp) add: token_lemma vimage_Diff HasTok_def) prefer 2 apply blast apply clarify apply (rule TR7_aux [THEN leadsTo_weaken]) apply (auto simp add: HasTok_def nodeOrder_def) done text\Misra's TR8: a hungry process eventually eats\ lemma token_progress: "j F \ ({s. token s < N} \ H j) leadsTo (E j)" apply (rule leadsTo_cancel1 [THEN leadsTo_Un_duplicate]) apply (rule_tac [2] TR6) apply (rule psp [OF leadsTo_j TR3, THEN leadsTo_weaken], blast+) done end end