diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,1361 +1,1359 @@ (* Title: HOL/Divides.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge *) section \More on quotient and remainder\ theory Divides imports Parity begin subsection \More on division\ inductive eucl_rel_int :: "int \ int \ int \ int \ bool" where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" | eucl_rel_int_dividesI: "l \ 0 \ k = q * l \ eucl_rel_int k l (q, 0)" | eucl_rel_int_remainderI: "sgn r = sgn l \ \r\ < \l\ \ k = q * l + r \ eucl_rel_int k l (q, r)" lemma eucl_rel_int_iff: "eucl_rel_int k l (q, r) \ k = l * q + r \ (if 0 < l then 0 \ r \ r < l else if l < 0 then l < r \ r \ 0 else q = 0)" by (cases "r = 0") (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI simp add: ac_simps sgn_1_pos sgn_1_neg) lemma unique_quotient_lemma: assumes "b * q' + r' \ b * q + r" "0 \ r'" "r' < b" "r < b" shows "q' \ (q::int)" proof - have "r' + b * (q'-q) \ r" using assms by (simp add: right_diff_distrib) moreover have "0 < b * (1 + q - q') " using assms by (simp add: right_diff_distrib distrib_left) moreover have "b * q' < b * (1 + q)" using assms by (simp add: right_diff_distrib distrib_left) ultimately show ?thesis using assms by (simp add: mult_less_cancel_left) qed lemma unique_quotient_lemma_neg: "b * q' + r' \ b*q + r \ r \ 0 \ b < r \ b < r' \ q \ (q'::int)" by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma) auto lemma unique_quotient: "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ q = q'" apply (rule order_antisym) apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ done lemma unique_remainder: "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ r = r'" apply (subgoal_tac "q = q'") apply (simp add: eucl_rel_int_iff) apply (blast intro: unique_quotient) done lemma eucl_rel_int: "eucl_rel_int k l (k div l, k mod l)" proof (cases k rule: int_cases3) case zero then show ?thesis by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def) next case (pos n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) next case (neg n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) qed lemma divmod_int_unique: assumes "eucl_rel_int k l (q, r)" shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" using assms eucl_rel_int [of k l] using unique_quotient [of k l] unique_remainder [of k l] by auto lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (simp add: divide_int_def) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma zdiv_int: "int (a div b) = int a div int b" by (simp add: divide_int_def) lemma zmod_int: "int (a mod b) = int a mod int b" by (simp add: modulo_int_def) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" proof - from assms have "sgn v = - 1 \ sgn v = 1" by (cases "v \ 0") auto then show ?thesis using assms unfolding divide_int_def [of "sgn v * \k\" "sgn v * \l\"] by (fastforce simp add: not_less div_abs_eq_div_nat) qed lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" proof (cases "l = 0") case True then show ?thesis by simp next case False with assms have "(sgn k * \k\) div (sgn l * \l\) = \k\ div \l\" using div_sgn_abs_cancel [of l k l] by simp then show ?thesis by (simp add: sgn_mult_abs) qed lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" proof (cases "k = 0 \ l = 0") case True then show ?thesis by auto next case False then have "k \ 0" and "l \ 0" by auto show ?thesis proof (cases "sgn l = sgn k") case True then show ?thesis by (simp add: div_eq_sgn_abs) next case False with \k \ 0\ \l \ 0\ have "sgn l * sgn k = - 1" by (simp add: sgn_if split: if_splits) with assms show ?thesis unfolding divide_int_def [of k l] by (auto simp add: zdiv_int ac_simps) qed qed 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 (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int) subsubsection \Laws for div and mod with Unary Minus\ lemma zminus1_lemma: "eucl_rel_int a b (q, r) ==> b \ 0 ==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1, if r=0 then 0 else b-r)" by (force simp add: eucl_rel_int_iff right_diff_distrib) lemma zdiv_zminus1_eq_if: "b \ (0::int) \ (-a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)" by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN div_int_unique]) lemma zmod_zminus1_eq_if: "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" proof (cases "b = 0") case False then show ?thesis by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN mod_int_unique]) qed auto 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_zminus2_eq_if: "b \ (0::int) ==> a div (-b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)" by (auto simp add: zdiv_zminus1_eq_if div_minus_right) lemma zmod_zminus2_eq_if: "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" by (auto simp add: zmod_zminus1_eq_if mod_minus_right) subsubsection \Monotonicity in the First Argument (Dividend)\ lemma zdiv_mono1: fixes b::int assumes "a \ a'" "0 < b" shows "a div b \ a' div b" proof (rule unique_quotient_lemma) show "b * (a div b) + a mod b \ b * (a' div b) + a' mod b" using assms(1) by auto qed (use assms in auto) lemma zdiv_mono1_neg: fixes b::int assumes "a \ a'" "b < 0" shows "a' div b \ a div b" proof (rule unique_quotient_lemma_neg) show "b * (a div b) + a mod b \ b * (a' div b) + a' mod b" using assms(1) by auto qed (use assms in auto) subsubsection \Monotonicity in the Second Argument (Divisor)\ lemma q_pos_lemma: fixes q'::int assumes "0 \ b'*q' + r'" "r' < b'" "0 < b'" shows "0 \ q'" proof - have "0 < b'* (q' + 1)" using assms by (simp add: distrib_left) with assms show ?thesis by (simp add: zero_less_mult_iff) qed lemma zdiv_mono2_lemma: fixes q'::int assumes eq: "b*q + r = b'*q' + r'" and le: "0 \ b'*q' + r'" and "r' < b'" "0 \ r" "0 < b'" "b' \ b" shows "q \ q'" proof - have "0 \ q'" using q_pos_lemma le \r' < b'\ \0 < b'\ by blast moreover have "b*q = r' - r + b'*q'" using eq by linarith ultimately have "b*q < b* (q' + 1)" using mult_right_mono assms unfolding distrib_left by fastforce with assms show ?thesis by (simp add: mult_less_cancel_left_pos) qed lemma zdiv_mono2: fixes a::int assumes "0 \ a" "0 < b'" "b' \ b" shows "a div b \ a div b'" proof (rule zdiv_mono2_lemma) have "b \ 0" using assms by linarith show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" by simp qed (use assms in auto) lemma zdiv_mono2_neg_lemma: fixes q'::int assumes "b*q + r = b'*q' + r'" "b'*q' + r' < 0" "r < b" "0 \ r'" "0 < b'" "b' \ b" shows "q' \ q" proof - have "b'*q' < 0" using assms by linarith with assms have "q' \ 0" by (simp add: mult_less_0_iff) have "b*q' \ b'*q'" by (simp add: \q' \ 0\ assms(6) mult_right_mono_neg) then have "b*q' < b* (q + 1)" using assms by (simp add: distrib_left) then show ?thesis using assms by (simp add: mult_less_cancel_left) qed lemma zdiv_mono2_neg: fixes a::int assumes "a < 0" "0 < b'" "b' \ b" shows "a div b' \ a div b" proof (rule zdiv_mono2_neg_lemma) have "b \ 0" using assms by linarith show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" by simp qed (use assms in auto) lemma div_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k div l = (k - l) div l + 1" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by (simp add: div_add_self2) qed lemma mod_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k mod l = (k - l) mod l" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by simp qed subsubsection \Splitting Rules for div and mod\ text\The proofs of the two lemmas below are essentially identical\ lemma split_pos_lemma: "0 P(n div k :: int)(n mod k) = (\i j. 0\j \ j n = k*i + j \ P i j)" by auto lemma split_neg_lemma: "k<0 \ P(n div k :: int)(n mod k) = (\i j. k j\0 \ n = k*i + j \ P i j)" by auto lemma split_zdiv: "P(n div k :: int) = ((k = 0 \ P 0) \ (0 (\i j. 0\j \ j n = k*i + j \ P i)) \ (k<0 \ (\i j. k j\0 \ n = k*i + j \ P i)))" proof (cases "k = 0") case False then show ?thesis unfolding linorder_neq_iff by (auto simp add: split_pos_lemma [of concl: "\x y. P x"] split_neg_lemma [of concl: "\x y. P x"]) qed auto lemma split_zmod: "P(n mod k :: int) = ((k = 0 \ P n) \ (0 (\i j. 0\j \ j n = k*i + j \ P j)) \ (k<0 \ (\i j. k j\0 \ n = k*i + j \ P j)))" proof (cases "k = 0") case False then show ?thesis unfolding linorder_neq_iff by (auto simp add: split_pos_lemma [of concl: "\x y. P y"] split_neg_lemma [of concl: "\x y. P y"]) qed auto 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 k", arith_split] for k declare split_zmod [of _ _ "numeral k", arith_split] for k subsubsection \Computing \div\ and \mod\ with shifting\ lemma pos_eucl_rel_int_mult_2: assumes "0 \ b" assumes "eucl_rel_int a b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)" using assms unfolding eucl_rel_int_iff by auto lemma neg_eucl_rel_int_mult_2: assumes "b \ 0" assumes "eucl_rel_int (a + 1) b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)" using assms unfolding eucl_rel_int_iff by auto text\computing div by shifting\ lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int] by (rule div_int_unique) lemma neg_zdiv_mult_2: assumes A: "a \ (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int] by (rule div_int_unique) 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 pos_zmod_mult_2: fixes a b :: int assumes "0 \ a" shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) lemma neg_zmod_mult_2: fixes a b :: int assumes "a \ 0" shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) 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) 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_trival_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" by safe (insert div_mult_mod_eq [of i k], auto) with zdiv_eq_0_iff show ?thesis by simp 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 \0 < k mod l\ \k mod l < 1 + l\ using \0 < l\ le_imp_0_less False apply auto using le_less apply fastforce using pos_mod_bound [of l k] apply linarith done with \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ by (simp add: zmod_trival_iff) ultimately show ?thesis apply (simp only: zmod_zminus1_eq_if) apply (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) done qed qed lemma div_neg_pos_less0: fixes a::int assumes "a < 0" "0 < b" shows "a div b < 0" proof - have "a div b \ - 1 div b" using zdiv_mono1 assms by auto also have "... \ -1" by (simp add: assms(2) div_eq_minus1) finally show ?thesis by force qed lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" by (drule zdiv_mono1_neg, auto) lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" by (drule zdiv_mono1, auto) 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: fixes a::int assumes "0 < b" shows "(0 \ a div b) = (0 \ a)" proof show "0 \ a div b \ 0 \ a" using assms by (simp add: linorder_not_less [symmetric]) (blast intro: div_neg_pos_less0) next assume "0 \ a" then have "0 div b \ a div b" using zdiv_mono1 assms by blast then show "0 \ a div b" by auto qed lemma pos_imp_zdiv_pos_iff: "0 0 < (i::int) div k \ k \ i" using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] by arith lemma neg_imp_zdiv_nonneg_iff: fixes a::int assumes "b < 0" shows "(0 \ a div b) = (a \ 0)" using assms by (simp add: div_minus_minus [of a, symmetric] pos_imp_zdiv_nonneg_iff del: div_minus_minus) (*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) (*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) lemma nonneg1_imp_zdiv_pos_iff: fixes a::int assumes "0 \ a" shows "a div b > 0 \ a \ b \ b>0" proof - have "0 < a div b \ b \ a" using div_pos_pos_trivial[of a b] assms by arith moreover have "0 < a div b \ b > 0" using assms 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::int) \ 0 \ m mod k \ m" by (rule split_zmod[THEN iffD2]) (fastforce dest: q_pos_lemma intro: split_mult_pos_le) 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 show ?thesis apply (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) 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 text \Simplify expressions in which div and mod combine numerical constants\ lemma int_div_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a div b = q" by (rule div_int_unique [of a b q r]) (simp add: eucl_rel_int_iff) lemma int_div_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a div b = q" by (rule div_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma abs_div: "(y::int) dvd x \ \x div y\ = \x\ div \y\" unfolding dvd_def by (cases "y=0") (auto simp add: abs_mult) text\Suggested by Matthias Daum\ lemma int_power_div_base: fixes k :: int assumes "0 < m" "0 < k" shows "k ^ m div k = (k::int) ^ (m - Suc 0)" proof - have eq: "k ^ m = k ^ ((m - Suc 0) + Suc 0)" by (simp add: assms) show ?thesis using assms by (simp only: power_add eq) auto qed text\Suggested by Matthias Daum\ lemma int_div_less_self: fixes x::int assumes "0 < x" "1 < k" shows "x div k < x" proof - have "nat x div nat k < nat x" by (simp add: assms) with assms show ?thesis by (simp add: nat_div_distrib [symmetric]) qed lemma mod_eq_dvd_iff_nat: "m mod q = n mod q \ q dvd m - n" if "m \ n" for m n q :: nat proof - have "int m mod int q = int n mod int q \ int q dvd int m - int n" by (simp add: mod_eq_dvd_iff) with that have "int (m mod q) = int (n mod q) \ int q dvd int (m - n)" by (simp only: of_nat_mod of_nat_diff) then show ?thesis by simp qed 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_lemma: assumes "(x::nat) mod n = y mod n" and "y \ x" shows "\q. x = y + n * q" using assms by (rule mod_eq_nat1E) rule 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 nat_mod_eq_lemma[OF th xy] have ?rhs apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} moreover {assume xy: "y \ x" from nat_mod_eq_lemma[OF H xy] have ?rhs apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} 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 -lemma take_bit_greater_eq: - \k + 2 ^ n \ take_bit n k\ if \k < 0\ for k :: int -proof - - have \k + 2 ^ n \ take_bit n (k + 2 ^ n)\ - proof (cases \k > - (2 ^ n)\) - case False - then have \k + 2 ^ n \ 0\ - by simp - also note take_bit_nonnegative - finally show ?thesis . - next - case True - with that have \0 \ k + 2 ^ n\ and \k + 2 ^ n < 2 ^ n\ - by simp_all - then show ?thesis - by (simp only: take_bit_eq_mod mod_pos_pos_trivial) - qed - then show ?thesis - by (simp add: take_bit_eq_mod) -qed - -lemma take_bit_less_eq: - \take_bit n k \ k - 2 ^ n\ if \2 ^ n \ k\ and \n > 0\ for k :: int - using that zmod_le_nonneg_dividend [of \k - 2 ^ n\ \2 ^ n\] - by (simp add: take_bit_eq_mod) - subsection \Numeral division with a pragmatic type class\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. This is its primary motivation, and it could surely be formulated using a more fine-grained, more algebraic and less technical class hierarchy. \ 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" fixes divmod :: "num \ num \ 'a \ 'a" and divmod_step :: "num \ 'a \ 'a \ 'a \ 'a" assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)" and divmod_step_def: "divmod_step l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" \ \These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ begin lemma divmod_digit_1: 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" by (auto simp add: mod_w) (insert mod_less, auto) 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: 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: 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 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 \ 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. \ lemma divmod_step_eq [simp]: "divmod_step l (q, r) = (if numeral l \ r then (2 * q + 1, r - numeral l) else (2 * q, r))" by (simp add: divmod_step_def) text \ This is 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 n (divmod m (Num.Bit0 n)))" proof (cases "m < n") case True then have "numeral m < numeral n" by simp then show ?thesis by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod) next case False have "divmod m n = divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n))" proof (cases "numeral n \ numeral m mod (2 * numeral n)") case True with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)" by simp moreover from True divmod_digit_1 [of "numeral m" "numeral n"] have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n" and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n" by simp_all ultimately show ?thesis by (simp only: divmod_def) next case False then have *: "numeral m mod (2 * numeral n) < numeral n" by (simp add: not_le) with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))" by auto moreover from * divmod_digit_0 [of "numeral n" "numeral m"] have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n" and "numeral m mod (2 * numeral n) = numeral m mod numeral n" by (simp_all only: zero_less_numeral) ultimately show ?thesis by (simp only: divmod_def) qed then have "divmod m n = divmod_step n (numeral m div numeral (Num.Bit0 n), numeral m mod numeral (Num.Bit0 n))" by (simp only: numeral.simps distrib mult_1) then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))" by (simp add: divmod_def) with False show ?thesis by simp 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 - have *: "\q. numeral (Num.Bit0 q) = 2 * numeral q" "\q. numeral (Num.Bit1 q) = 2 * numeral q + 1" by (simp_all only: numeral_mult numeral.simps distrib) simp_all have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less) then show ?P and ?Q by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1 div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] add.commute del: numeral_times_numeral) 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 (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 (num.Bit1 n) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_step_eq 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) text \Computing congruences modulo \2 ^ q\\ lemma cong_exp_iff_simps: "numeral n mod numeral Num.One = 0 \ True" "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 \ numeral n mod numeral q = 0" "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 \ False" "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ (numeral n mod numeral q) = 0" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ (numeral m mod numeral q) = 0" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) end hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq instantiation nat :: unique_euclidean_semiring_numeral 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 :: "num \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: fixes k l :: num shows "Suc 0 div numeral k = fst (divmod Num.One k)" by (simp_all add: fst_divmod) lemma Suc_0_mod_numeral [simp]: fixes k l :: num shows "Suc 0 mod numeral k = snd (divmod Num.One k)" by (simp_all add: snd_divmod) instantiation int :: unique_euclidean_semiring_numeral 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 :: "num \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) 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 :: "int \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else 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 (numeral 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 (numeral 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 (numeral 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 (numeral n) (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp 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 lemma bit_numeral_Bit0_Suc_iff [simp]: \bit (numeral (Num.Bit0 m) :: int) (Suc n) \ bit (numeral m :: int) n\ by (simp add: bit_Suc) lemma bit_numeral_Bit1_Suc_iff [simp]: \bit (numeral (Num.Bit1 m) :: int) (Suc n) \ bit (numeral m :: int) n\ by (simp add: bit_Suc) lemma div_positive_int: "k div l > 0" if "k \ l" and "l > 0" for k l :: int using that div_positive [of l k] by blast subsubsection \Dedicated simproc for calculation\ text \ 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. \ simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "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_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "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 Divides.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_eq fst_conv snd_conv numeral_One case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right} @ [@{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 \ subsubsection \Code generation\ definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" lemma fst_divmod_nat [simp]: "fst (divmod_nat m n) = m div n" by (simp add: divmod_nat_def) lemma snd_divmod_nat [simp]: "snd (divmod_nat m n) = m mod n" by (simp add: divmod_nat_def) lemma divmod_nat_if [code]: "Divides.divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))" by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by simp_all 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 = - (Divides.adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - Divides.adjust_mod (Int.Pos 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 code_identifier code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \More on bit operations\ lemma take_bit_incr_eq: \take_bit n (k + 1) = 1 + take_bit n k\ if \take_bit n k \ 2 ^ n - 1\ for k :: int proof - from that have \2 ^ n \ k mod 2 ^ n + 1\ by (simp add: take_bit_eq_mod) moreover have \k mod 2 ^ n < 2 ^ n\ by simp ultimately have *: \k mod 2 ^ n + 1 < 2 ^ n\ by linarith have \(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\ by (simp add: mod_simps) also have \\ = k mod 2 ^ n + 1\ using * by (simp add: zmod_trival_iff) finally have \(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\ . then show ?thesis by (simp add: take_bit_eq_mod) qed lemma take_bit_decr_eq: \take_bit n (k - 1) = take_bit n k - 1\ if \take_bit n k \ 0\ for k :: int proof - from that have \k mod 2 ^ n \ 0\ by (simp add: take_bit_eq_mod) moreover have \k mod 2 ^ n \ 0\ \k mod 2 ^ n < 2 ^ n\ by simp_all ultimately have *: \k mod 2 ^ n > 0\ by linarith have \(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\ by (simp add: mod_simps) also have \\ = k mod 2 ^ n - 1\ by (simp add: zmod_trival_iff) (use \k mod 2 ^ n < 2 ^ n\ * in linarith) finally have \(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\ . then show ?thesis by (simp add: take_bit_eq_mod) qed +lemma take_bit_int_greater_eq: + \k + 2 ^ n \ take_bit n k\ if \k < 0\ for k :: int +proof - + have \k + 2 ^ n \ take_bit n (k + 2 ^ n)\ + proof (cases \k > - (2 ^ n)\) + case False + then have \k + 2 ^ n \ 0\ + by simp + also note take_bit_nonnegative + finally show ?thesis . + next + case True + with that have \0 \ k + 2 ^ n\ and \k + 2 ^ n < 2 ^ n\ + by simp_all + then show ?thesis + by (simp only: take_bit_eq_mod mod_pos_pos_trivial) + qed + then show ?thesis + by (simp add: take_bit_eq_mod) +qed + +lemma take_bit_int_less_eq: + \take_bit n k \ k - 2 ^ n\ if \2 ^ n \ k\ and \n > 0\ for k :: int + using that zmod_le_nonneg_dividend [of \k - 2 ^ n\ \2 ^ n\] + by (simp add: take_bit_eq_mod) + lemma take_bit_int_less_eq_self_iff: \take_bit n k \ k \ 0 \ k\ (is \?P \ ?Q\) for k :: int proof assume ?P show ?Q proof (rule ccontr) assume \\ 0 \ k\ then have \k < 0\ by simp with \?P\ have \take_bit n k < 0\ by (rule le_less_trans) then show False by simp qed next assume ?Q then show ?P by (simp add: take_bit_eq_mod zmod_le_nonneg_dividend) qed lemma take_bit_int_less_self_iff: \take_bit n k < k \ 2 ^ n \ k\ for k :: int by (auto simp add: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff intro: order_trans [of 0 \2 ^ n\ k]) lemma take_bit_int_greater_self_iff: \k < take_bit n k \ k < 0\ for k :: int using take_bit_int_less_eq_self_iff [of n k] by auto lemma take_bit_int_greater_eq_self_iff: \k \ take_bit n k \ k < 2 ^ n\ for k :: int by (auto simp add: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff dest: sym not_sym intro: less_trans [of k 0 \2 ^ n\]) subsection \Lemmas of doubtful value\ 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 -find_theorems \(?k::int) mod _ = ?k\ - end diff --git a/src/HOL/Library/Bit_Operations.thy b/src/HOL/Library/Bit_Operations.thy --- a/src/HOL/Library/Bit_Operations.thy +++ b/src/HOL/Library/Bit_Operations.thy @@ -1,1421 +1,1440 @@ (* Author: Florian Haftmann, TUM *) section \Bit operations in suitable algebraic structures\ theory Bit_Operations imports "HOL-Library.Boolean_Algebra" Main begin subsection \Bit operations\ class semiring_bit_operations = semiring_bit_shifts + fixes "and" :: \'a \ 'a \ 'a\ (infixr \AND\ 64) and or :: \'a \ 'a \ 'a\ (infixr \OR\ 59) and xor :: \'a \ 'a \ 'a\ (infixr \XOR\ 59) and mask :: \nat \ 'a\ assumes bit_and_iff: \\n. bit (a AND b) n \ bit a n \ bit b n\ and bit_or_iff: \\n. bit (a OR b) n \ bit a n \ bit b n\ and bit_xor_iff: \\n. bit (a XOR b) n \ bit a n \ bit b n\ and mask_eq_exp_minus_1: \mask n = 2 ^ n - 1\ begin text \ We want the bitwise operations to bind slightly weaker than \+\ and \-\. For the sake of code generation the operations \<^const>\and\, \<^const>\or\ and \<^const>\xor\ are specified as definitional class operations. \ sublocale "and": semilattice \(AND)\ by standard (auto simp add: bit_eq_iff bit_and_iff) sublocale or: semilattice_neutr \(OR)\ 0 by standard (auto simp add: bit_eq_iff bit_or_iff) sublocale xor: comm_monoid \(XOR)\ 0 by standard (auto simp add: bit_eq_iff bit_xor_iff) lemma even_and_iff: \even (a AND b) \ even a \ even b\ using bit_and_iff [of a b 0] by auto lemma even_or_iff: \even (a OR b) \ even a \ even b\ using bit_or_iff [of a b 0] by auto lemma even_xor_iff: \even (a XOR b) \ (even a \ even b)\ using bit_xor_iff [of a b 0] by auto lemma zero_and_eq [simp]: "0 AND a = 0" by (simp add: bit_eq_iff bit_and_iff) lemma and_zero_eq [simp]: "a AND 0 = 0" by (simp add: bit_eq_iff bit_and_iff) lemma one_and_eq: "1 AND a = a mod 2" by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff) lemma and_one_eq: "a AND 1 = a mod 2" using one_and_eq [of a] by (simp add: ac_simps) lemma one_or_eq: "1 OR a = a + of_bool (even a)" by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff) lemma or_one_eq: "a OR 1 = a + of_bool (even a)" using one_or_eq [of a] by (simp add: ac_simps) lemma one_xor_eq: "1 XOR a = a + of_bool (even a) - of_bool (odd a)" by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE) lemma xor_one_eq: "a XOR 1 = a + of_bool (even a) - of_bool (odd a)" using one_xor_eq [of a] by (simp add: ac_simps) lemma take_bit_and [simp]: \take_bit n (a AND b) = take_bit n a AND take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff) lemma take_bit_or [simp]: \take_bit n (a OR b) = take_bit n a OR take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff) lemma take_bit_xor [simp]: \take_bit n (a XOR b) = take_bit n a XOR take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff) lemma push_bit_and [simp]: \push_bit n (a AND b) = push_bit n a AND push_bit n b\ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff) lemma push_bit_or [simp]: \push_bit n (a OR b) = push_bit n a OR push_bit n b\ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff) lemma push_bit_xor [simp]: \push_bit n (a XOR b) = push_bit n a XOR push_bit n b\ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff) lemma drop_bit_and [simp]: \drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff) lemma drop_bit_or [simp]: \drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff) lemma drop_bit_xor [simp]: \drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff) lemma bit_mask_iff: \bit (mask m) n \ 2 ^ n \ 0 \ n < m\ by (simp add: mask_eq_exp_minus_1 bit_mask_iff) lemma even_mask_iff: \even (mask n) \ n = 0\ using bit_mask_iff [of n 0] by auto lemma mask_0 [simp]: \mask 0 = 0\ by (simp add: mask_eq_exp_minus_1) lemma mask_Suc_0 [simp]: \mask (Suc 0) = 1\ by (simp add: mask_eq_exp_minus_1 add_implies_diff sym) lemma mask_Suc_exp: \mask (Suc n) = 2 ^ n OR mask n\ by (rule bit_eqI) (auto simp add: bit_or_iff bit_mask_iff bit_exp_iff not_less le_less_Suc_eq) lemma mask_Suc_double: \mask (Suc n) = 1 OR 2 * mask n\ proof (rule bit_eqI) fix q assume \2 ^ q \ 0\ show \bit (mask (Suc n)) q \ bit (1 OR 2 * mask n) q\ by (cases q) (simp_all add: even_mask_iff even_or_iff bit_or_iff bit_mask_iff bit_exp_iff bit_double_iff not_less le_less_Suc_eq bit_1_iff, auto simp add: mult_2) qed lemma mask_numeral: \mask (numeral n) = 1 + 2 * mask (pred_numeral n)\ by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps) lemma take_bit_eq_mask: \take_bit n a = a AND mask n\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff) lemma disjunctive_add: \a + b = a OR b\ if \\n. \ bit a n \ \ bit b n\ by (rule bit_eqI) (use that in \simp add: bit_disjunctive_add_iff bit_or_iff\) end class ring_bit_operations = semiring_bit_operations + ring_parity + fixes not :: \'a \ 'a\ (\NOT\) assumes bit_not_iff: \\n. bit (NOT a) n \ 2 ^ n \ 0 \ \ bit a n\ assumes minus_eq_not_minus_1: \- a = NOT (a - 1)\ begin text \ For the sake of code generation \<^const>\not\ is specified as definitional class operation. Note that \<^const>\not\ has no sensible definition for unlimited but only positive bit strings (type \<^typ>\nat\). \ lemma bits_minus_1_mod_2_eq [simp]: \(- 1) mod 2 = 1\ by (simp add: mod_2_eq_odd) lemma not_eq_complement: \NOT a = - a - 1\ using minus_eq_not_minus_1 [of \a + 1\] by simp lemma minus_eq_not_plus_1: \- a = NOT a + 1\ using not_eq_complement [of a] by simp lemma bit_minus_iff: \bit (- a) n \ 2 ^ n \ 0 \ \ bit (a - 1) n\ by (simp add: minus_eq_not_minus_1 bit_not_iff) lemma even_not_iff [simp]: "even (NOT a) \ odd a" using bit_not_iff [of a 0] by auto lemma bit_not_exp_iff: \bit (NOT (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ by (auto simp add: bit_not_iff bit_exp_iff) lemma bit_minus_1_iff [simp]: \bit (- 1) n \ 2 ^ n \ 0\ by (simp add: bit_minus_iff) lemma bit_minus_exp_iff: \bit (- (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ oops lemma bit_minus_2_iff [simp]: \bit (- 2) n \ 2 ^ n \ 0 \ n > 0\ by (simp add: bit_minus_iff bit_1_iff) lemma not_one [simp]: "NOT 1 = - 2" by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) sublocale "and": semilattice_neutr \(AND)\ \- 1\ by standard (rule bit_eqI, simp add: bit_and_iff) sublocale bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ rewrites \bit.xor = (XOR)\ proof - interpret bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) show \boolean_algebra (AND) (OR) NOT 0 (- 1)\ by standard show \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ by (rule ext, rule ext, rule bit_eqI) (auto simp add: bit.xor_def bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) qed lemma and_eq_not_not_or: \a AND b = NOT (NOT a OR NOT b)\ by simp lemma or_eq_not_not_and: \a OR b = NOT (NOT a AND NOT b)\ by simp lemma not_add_distrib: \NOT (a + b) = NOT a - b\ by (simp add: not_eq_complement algebra_simps) lemma not_diff_distrib: \NOT (a - b) = NOT a + b\ using not_add_distrib [of a \- b\] by simp lemma disjunctive_diff: \a - b = a AND NOT b\ if \\n. bit b n \ bit a n\ proof - have \NOT a + b = NOT a OR b\ by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that) then have \NOT (NOT a + b) = NOT (NOT a OR b)\ by simp then show ?thesis by (simp add: not_add_distrib) qed lemma push_bit_minus: \push_bit n (- a) = - push_bit n a\ by (simp add: push_bit_eq_mult) lemma take_bit_not_take_bit: \take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff) lemma take_bit_not_iff: "take_bit n (NOT a) = take_bit n (NOT b) \ take_bit n a = take_bit n b" apply (simp add: bit_eq_iff) apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff) apply (use exp_eq_0_imp_not_bit in blast) done +lemma take_bit_not_eq_mask_diff: + \take_bit n (NOT a) = mask n - take_bit n a\ +proof - + have \take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\ + by (simp add: take_bit_not_take_bit) + also have \\ = mask n AND NOT (take_bit n a)\ + by (simp add: take_bit_eq_mask ac_simps) + also have \\ = mask n - take_bit n a\ + by (subst disjunctive_diff) + (auto simp add: bit_take_bit_iff bit_mask_iff exp_eq_0_imp_not_bit) + finally show ?thesis + by simp +qed + lemma mask_eq_take_bit_minus_one: \mask n = take_bit n (- 1)\ by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute) lemma take_bit_minus_one_eq_mask: \take_bit n (- 1) = mask n\ by (simp add: mask_eq_take_bit_minus_one) lemma minus_exp_eq_not_mask: \- (2 ^ n) = NOT (mask n)\ by (rule bit_eqI) (simp add: bit_minus_iff bit_not_iff flip: mask_eq_exp_minus_1) lemma push_bit_minus_one_eq_not_mask: \push_bit n (- 1) = NOT (mask n)\ by (simp add: push_bit_eq_mult minus_exp_eq_not_mask) lemma take_bit_not_mask_eq_0: \take_bit m (NOT (mask n)) = 0\ if \n \ m\ by (rule bit_eqI) (use that in \simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\) lemma take_bit_mask [simp]: \take_bit m (mask n) = mask (min m n)\ by (simp add: mask_eq_take_bit_minus_one) definition set_bit :: \nat \ 'a \ 'a\ where \set_bit n a = a OR push_bit n 1\ definition unset_bit :: \nat \ 'a \ 'a\ where \unset_bit n a = a AND NOT (push_bit n 1)\ definition flip_bit :: \nat \ 'a \ 'a\ where \flip_bit n a = a XOR push_bit n 1\ lemma bit_set_bit_iff: \bit (set_bit m a) n \ bit a n \ (m = n \ 2 ^ n \ 0)\ by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff) lemma even_set_bit_iff: \even (set_bit m a) \ even a \ m \ 0\ using bit_set_bit_iff [of m a 0] by auto lemma bit_unset_bit_iff: \bit (unset_bit m a) n \ bit a n \ m \ n\ by (auto simp add: unset_bit_def push_bit_of_1 bit_and_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit) lemma even_unset_bit_iff: \even (unset_bit m a) \ even a \ m = 0\ using bit_unset_bit_iff [of m a 0] by auto lemma bit_flip_bit_iff: \bit (flip_bit m a) n \ (m = n \ \ bit a n) \ 2 ^ n \ 0\ by (auto simp add: flip_bit_def push_bit_of_1 bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit) lemma even_flip_bit_iff: \even (flip_bit m a) \ \ (even a \ m = 0)\ using bit_flip_bit_iff [of m a 0] by auto lemma set_bit_0 [simp]: \set_bit 0 a = 1 + 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\ by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff) (cases m, simp_all add: bit_Suc) qed lemma set_bit_Suc: \set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ show \bit (set_bit (Suc n) a) m \ bit (a mod 2 + 2 * set_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_set_bit_iff) next case (Suc m) with * have \2 ^ m \ 0\ using mult_2 by auto show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *, simp_all add: Suc \2 ^ m \ 0\ bit_Suc) qed qed lemma unset_bit_0 [simp]: \unset_bit 0 a = 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\ by (simp add: bit_unset_bit_iff bit_double_iff) (cases m, simp_all add: bit_Suc) qed lemma unset_bit_Suc: \unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (unset_bit (Suc n) a) m \ bit (a mod 2 + 2 * unset_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_unset_bit_iff) next case (Suc m) show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *, simp_all add: Suc bit_Suc) qed qed lemma flip_bit_0 [simp]: \flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\ by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff) (cases m, simp_all add: bit_Suc) qed lemma flip_bit_Suc: \flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ show \bit (flip_bit (Suc n) a) m \ bit (a mod 2 + 2 * flip_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_flip_bit_iff) next case (Suc m) with * have \2 ^ m \ 0\ using mult_2 by auto show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff, simp_all add: Suc \2 ^ m \ 0\ bit_Suc) qed qed lemma flip_bit_eq_if: \flip_bit n a = (if bit a n then unset_bit else set_bit) n a\ by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff) lemma take_bit_set_bit_eq: \take_bit n (set_bit m a) = (if n \ m then take_bit n a else set_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff) lemma take_bit_unset_bit_eq: \take_bit n (unset_bit m a) = (if n \ m then take_bit n a else unset_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff) lemma take_bit_flip_bit_eq: \take_bit n (flip_bit m a) = (if n \ m then take_bit n a else flip_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff) end subsection \Instance \<^typ>\int\\ instantiation int :: ring_bit_operations begin definition not_int :: \int \ int\ where \not_int k = - k - 1\ lemma not_int_rec: "NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int by (auto simp add: not_int_def elim: oddE) lemma even_not_iff_int: \even (NOT k) \ odd k\ for k :: int by (simp add: not_int_def) lemma not_int_div_2: \NOT k div 2 = NOT (k div 2)\ for k :: int by (simp add: not_int_def) lemma bit_not_int_iff: \bit (NOT k) n \ \ bit k n\ for k :: int by (induction n arbitrary: k) (simp_all add: not_int_div_2 even_not_iff_int bit_Suc) function and_int :: \int \ int \ int\ where \(k::int) AND l = (if k \ {0, - 1} \ l \ {0, - 1} then - of_bool (odd k \ odd l) else of_bool (odd k \ odd l) + 2 * ((k div 2) AND (l div 2)))\ by auto termination by (relation \measure (\(k, l). nat (\k\ + \l\))\) auto declare and_int.simps [simp del] lemma and_int_rec: \k AND l = of_bool (odd k \ odd l) + 2 * ((k div 2) AND (l div 2))\ for k l :: int proof (cases \k \ {0, - 1} \ l \ {0, - 1}\) case True then show ?thesis by auto (simp_all add: and_int.simps) next case False then show ?thesis by (auto simp add: ac_simps and_int.simps [of k l]) qed lemma bit_and_int_iff: \bit (k AND l) n \ bit k n \ bit l n\ for k l :: int proof (induction n arbitrary: k l) case 0 then show ?case by (simp add: and_int_rec [of k l]) next case (Suc n) then show ?case by (simp add: and_int_rec [of k l] bit_Suc) qed lemma even_and_iff_int: \even (k AND l) \ even k \ even l\ for k l :: int using bit_and_int_iff [of k l 0] by auto definition or_int :: \int \ int \ int\ where \k OR l = NOT (NOT k AND NOT l)\ for k l :: int lemma or_int_rec: \k OR l = of_bool (odd k \ odd l) + 2 * ((k div 2) OR (l div 2))\ for k l :: int using and_int_rec [of \NOT k\ \NOT l\] by (simp add: or_int_def even_not_iff_int not_int_div_2) (simp add: not_int_def) lemma bit_or_int_iff: \bit (k OR l) n \ bit k n \ bit l n\ for k l :: int by (simp add: or_int_def bit_not_int_iff bit_and_int_iff) definition xor_int :: \int \ int \ int\ where \k XOR l = k AND NOT l OR NOT k AND l\ for k l :: int lemma xor_int_rec: \k XOR l = of_bool (odd k \ odd l) + 2 * ((k div 2) XOR (l div 2))\ for k l :: int by (simp add: xor_int_def or_int_rec [of \k AND NOT l\ \NOT k AND l\] even_and_iff_int even_not_iff_int) (simp add: and_int_rec [of \NOT k\ \l\] and_int_rec [of \k\ \NOT l\] not_int_div_2) lemma bit_xor_int_iff: \bit (k XOR l) n \ bit k n \ bit l n\ for k l :: int by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff) definition mask_int :: \nat \ int\ where \mask n = (2 :: int) ^ n - 1\ instance proof fix k l :: int and n :: nat show \- k = NOT (k - 1)\ by (simp add: not_int_def) show \bit (k AND l) n \ bit k n \ bit l n\ by (fact bit_and_int_iff) show \bit (k OR l) n \ bit k n \ bit l n\ by (fact bit_or_int_iff) show \bit (k XOR l) n \ bit k n \ bit l n\ by (fact bit_xor_int_iff) qed (simp_all add: bit_not_int_iff mask_int_def) end lemma mask_half_int: \mask n div 2 = (mask (n - 1) :: int)\ by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps) lemma mask_nonnegative_int [simp]: \mask n \ (0::int)\ by (simp add: mask_eq_exp_minus_1) lemma not_mask_negative_int [simp]: \\ mask n < (0::int)\ by (simp add: not_less) lemma not_nonnegative_int_iff [simp]: \NOT k \ 0 \ k < 0\ for k :: int by (simp add: not_int_def) lemma not_negative_int_iff [simp]: \NOT k < 0 \ k \ 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le) lemma and_nonnegative_int_iff [simp]: \k AND l \ 0 \ k \ 0 \ l \ 0\ for k l :: int proof (induction k arbitrary: l rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) then show ?case using and_int_rec [of \k * 2\ l] by (simp add: pos_imp_zdiv_nonneg_iff) next case (odd k) from odd have \0 \ k AND l div 2 \ 0 \ k \ 0 \ l div 2\ by simp then have \0 \ (1 + k * 2) div 2 AND l div 2 \ 0 \ (1 + k * 2) div 2\ 0 \ l div 2\ by simp with and_int_rec [of \1 + k * 2\ l] show ?case by auto qed lemma and_negative_int_iff [simp]: \k AND l < 0 \ k < 0 \ l < 0\ for k l :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma and_less_eq: \k AND l \ k\ if \l < 0\ for k l :: int using that proof (induction k arbitrary: l rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) from even.IH [of \l div 2\] even.hyps even.prems show ?case by (simp add: and_int_rec [of _ l]) next case (odd k) from odd.IH [of \l div 2\] odd.hyps odd.prems show ?case by (simp add: and_int_rec [of _ l]) qed lemma or_nonnegative_int_iff [simp]: \k OR l \ 0 \ k \ 0 \ l \ 0\ for k l :: int by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp lemma or_negative_int_iff [simp]: \k OR l < 0 \ k < 0 \ l < 0\ for k l :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma or_greater_eq: \k OR l \ k\ if \l \ 0\ for k l :: int using that proof (induction k arbitrary: l rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) from even.IH [of \l div 2\] even.hyps even.prems show ?case by (simp add: or_int_rec [of _ l]) next case (odd k) from odd.IH [of \l div 2\] odd.hyps odd.prems show ?case by (simp add: or_int_rec [of _ l]) qed lemma xor_nonnegative_int_iff [simp]: \k XOR l \ 0 \ (k \ 0 \ l \ 0)\ for k l :: int by (simp only: bit.xor_def or_nonnegative_int_iff) auto lemma xor_negative_int_iff [simp]: \k XOR l < 0 \ (k < 0) \ (l < 0)\ for k l :: int by (subst Not_eq_iff [symmetric]) (auto simp add: not_less) lemma set_bit_nonnegative_int_iff [simp]: \set_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: set_bit_def) lemma set_bit_negative_int_iff [simp]: \set_bit n k < 0 \ k < 0\ for k :: int by (simp add: set_bit_def) lemma unset_bit_nonnegative_int_iff [simp]: \unset_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: unset_bit_def) lemma unset_bit_negative_int_iff [simp]: \unset_bit n k < 0 \ k < 0\ for k :: int by (simp add: unset_bit_def) lemma flip_bit_nonnegative_int_iff [simp]: \flip_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: flip_bit_def) lemma flip_bit_negative_int_iff [simp]: \flip_bit n k < 0 \ k < 0\ for k :: int by (simp add: flip_bit_def) lemma set_bit_greater_eq: \set_bit n k \ k\ for k :: int by (simp add: set_bit_def or_greater_eq) lemma unset_bit_less_eq: \unset_bit n k \ k\ for k :: int by (simp add: unset_bit_def and_less_eq) lemma set_bit_eq: \set_bit n k = k + of_bool (\ bit k n) * 2 ^ n\ for k :: int proof (rule bit_eqI) fix m show \bit (set_bit n k) m \ bit (k + of_bool (\ bit k n) * 2 ^ n) m\ proof (cases \m = n\) case True then show ?thesis apply (simp add: bit_set_bit_iff) apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right) done next case False then show ?thesis apply (clarsimp simp add: bit_set_bit_iff) apply (subst disjunctive_add) apply (clarsimp simp add: bit_exp_iff) apply (clarsimp simp add: bit_or_iff bit_exp_iff) done qed qed lemma unset_bit_eq: \unset_bit n k = k - of_bool (bit k n) * 2 ^ n\ for k :: int proof (rule bit_eqI) fix m show \bit (unset_bit n k) m \ bit (k - of_bool (bit k n) * 2 ^ n) m\ proof (cases \m = n\) case True then show ?thesis apply (simp add: bit_unset_bit_iff) apply (simp add: bit_iff_odd) using div_plus_div_distrib_dvd_right [of \2 ^ n\ \- (2 ^ n)\ k] apply (simp add: dvd_neg_div) done next case False then show ?thesis apply (clarsimp simp add: bit_unset_bit_iff) apply (subst disjunctive_diff) apply (clarsimp simp add: bit_exp_iff) apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff) done qed qed context ring_bit_operations begin lemma even_of_int_iff: \even (of_int k) \ even k\ by (induction k rule: int_bit_induct) simp_all lemma bit_of_int_iff: \bit (of_int k) n \ (2::'a) ^ n \ 0 \ bit k n\ proof (cases \(2::'a) ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False then have \bit (of_int k) n \ bit k n\ proof (induction k arbitrary: n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) then show ?case using bit_double_iff [of \of_int k\ n] Parity.bit_double_iff [of k n] by (cases n) (auto simp add: ac_simps dest: mult_not_zero) next case (odd k) then show ?case using bit_double_iff [of \of_int k\ n] by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero) qed with False show ?thesis by simp qed lemma push_bit_of_int: \push_bit n (of_int k) = of_int (push_bit n k)\ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) lemma of_int_push_bit: \of_int (push_bit n k) = push_bit n (of_int k)\ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) lemma take_bit_of_int: \take_bit n (of_int k) = of_int (take_bit n k)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff) lemma of_int_take_bit: \of_int (take_bit n k) = take_bit n (of_int k)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff) lemma of_int_not_eq: \of_int (NOT k) = NOT (of_int k)\ by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff) lemma of_int_and_eq: \of_int (k AND l) = of_int k AND of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff) lemma of_int_or_eq: \of_int (k OR l) = of_int k OR of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff) lemma of_int_xor_eq: \of_int (k XOR l) = of_int k XOR of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff) lemma of_int_mask_eq: \of_int (mask n) = mask n\ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq) end subsection \Bit concatenation\ definition concat_bit :: \nat \ int \ int \ int\ where \concat_bit n k l = take_bit n k OR push_bit n l\ lemma bit_concat_bit_iff: \bit (concat_bit m k l) n \ n < m \ bit k n \ m \ n \ bit l (n - m)\ by (simp add: concat_bit_def bit_or_iff bit_and_iff bit_take_bit_iff bit_push_bit_iff ac_simps) lemma concat_bit_eq: \concat_bit n k l = take_bit n k + push_bit n l\ by (simp add: concat_bit_def take_bit_eq_mask bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add) lemma concat_bit_0 [simp]: \concat_bit 0 k l = l\ by (simp add: concat_bit_def) lemma concat_bit_Suc: \concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\ by (simp add: concat_bit_eq take_bit_Suc push_bit_double) lemma concat_bit_of_zero_1 [simp]: \concat_bit n 0 l = push_bit n l\ by (simp add: concat_bit_def) lemma concat_bit_of_zero_2 [simp]: \concat_bit n k 0 = take_bit n k\ by (simp add: concat_bit_def take_bit_eq_mask) lemma concat_bit_nonnegative_iff [simp]: \concat_bit n k l \ 0 \ l \ 0\ by (simp add: concat_bit_def) lemma concat_bit_negative_iff [simp]: \concat_bit n k l < 0 \ l < 0\ by (simp add: concat_bit_def) lemma concat_bit_assoc: \concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps) lemma concat_bit_assoc_sym: \concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def) lemma concat_bit_eq_iff: \concat_bit n k l = concat_bit n r s \ take_bit n k = take_bit n r \ l = s\ (is \?P \ ?Q\) proof assume ?Q then show ?P by (simp add: concat_bit_def) next assume ?P then have *: \bit (concat_bit n k l) m = bit (concat_bit n r s) m\ for m by (simp add: bit_eq_iff) have \take_bit n k = take_bit n r\ proof (rule bit_eqI) fix m from * [of m] show \bit (take_bit n k) m \ bit (take_bit n r) m\ by (auto simp add: bit_take_bit_iff bit_concat_bit_iff) qed moreover have \push_bit n l = push_bit n s\ proof (rule bit_eqI) fix m from * [of m] show \bit (push_bit n l) m \ bit (push_bit n s) m\ by (auto simp add: bit_push_bit_iff bit_concat_bit_iff) qed then have \l = s\ by (simp add: push_bit_eq_mult) ultimately show ?Q by (simp add: concat_bit_def) qed lemma take_bit_concat_bit_eq: \take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def) subsection \Taking bits with sign propagation\ context ring_bit_operations begin definition signed_take_bit :: \nat \ 'a \ 'a\ where \signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\ lemma signed_take_bit_eq_if_positive: \signed_take_bit n a = take_bit n a\ if \\ bit a n\ using that by (simp add: signed_take_bit_def) lemma signed_take_bit_eq_if_negative: \signed_take_bit n a = take_bit n a OR NOT (mask n)\ if \bit a n\ using that by (simp add: signed_take_bit_def) lemma even_signed_take_bit_iff: \even (signed_take_bit m a) \ even a\ by (auto simp add: signed_take_bit_def even_or_iff even_mask_iff bit_double_iff) lemma bit_signed_take_bit_iff: \bit (signed_take_bit m a) n \ 2 ^ n \ 0 \ bit a (min m n)\ by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le) (use exp_eq_0_imp_not_bit in blast) lemma signed_take_bit_0 [simp]: \signed_take_bit 0 a = - (a mod 2)\ by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one) lemma signed_take_bit_Suc: \signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ show \bit (signed_take_bit (Suc n) a) m \ bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_signed_take_bit_iff) next case (Suc m) with * have \2 ^ m \ 0\ by (metis mult_not_zero power_Suc) with Suc show ?thesis by (simp add: bit_signed_take_bit_iff mod2_eq_if bit_double_iff even_bit_succ_iff ac_simps flip: bit_Suc) qed qed lemma signed_take_bit_of_0 [simp]: \signed_take_bit n 0 = 0\ by (simp add: signed_take_bit_def) lemma signed_take_bit_of_minus_1 [simp]: \signed_take_bit n (- 1) = - 1\ by (simp add: signed_take_bit_def take_bit_minus_one_eq_mask mask_eq_exp_minus_1) lemma signed_take_bit_Suc_1 [simp]: \signed_take_bit (Suc n) 1 = 1\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_rec: \signed_take_bit n a = (if n = 0 then - (a mod 2) else a mod 2 + 2 * signed_take_bit (n - 1) (a div 2))\ by (cases n) (simp_all add: signed_take_bit_Suc) lemma signed_take_bit_eq_iff_take_bit_eq: \signed_take_bit n a = signed_take_bit n b \ take_bit (Suc n) a = take_bit (Suc n) b\ proof - have \bit (signed_take_bit n a) = bit (signed_take_bit n b) \ bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\ by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def) (use exp_eq_0_imp_not_bit in fastforce) then show ?thesis by (simp add: bit_eq_iff fun_eq_iff) qed lemma signed_take_bit_signed_take_bit [simp]: \signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\ proof (rule bit_eqI) fix q show \bit (signed_take_bit m (signed_take_bit n a)) q \ bit (signed_take_bit (min m n) a) q\ by (simp add: bit_signed_take_bit_iff min_def bit_or_iff bit_not_iff bit_mask_iff bit_take_bit_iff) (use le_Suc_ex exp_add_not_zero_imp in blast) qed lemma signed_take_bit_take_bit: \signed_take_bit m (take_bit n a) = (if n \ m then take_bit n else signed_take_bit m) a\ by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff) lemma take_bit_signed_take_bit: \take_bit m (signed_take_bit n a) = take_bit m a\ if \m \ Suc n\ using that by (rule le_SucE; intro bit_eqI) (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq) end text \Modulus centered around 0\ lemma signed_take_bit_eq_concat_bit: \signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\ by (simp add: concat_bit_def signed_take_bit_def push_bit_minus_one_eq_not_mask) lemma signed_take_bit_add: \signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) + take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k + l)\ by (simp add: take_bit_signed_take_bit take_bit_add) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add) qed lemma signed_take_bit_diff: \signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) - take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k - l)\ by (simp add: take_bit_signed_take_bit take_bit_diff) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff) qed lemma signed_take_bit_minus: \signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\ for k :: int proof - have \take_bit (Suc n) (- take_bit (Suc n) (signed_take_bit n k)) = take_bit (Suc n) (- k)\ by (simp add: take_bit_signed_take_bit take_bit_minus) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus) qed lemma signed_take_bit_mult: \signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) * take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k * l)\ by (simp add: take_bit_signed_take_bit take_bit_mult) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult) qed lemma signed_take_bit_eq_take_bit_minus: \signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\ for k :: int proof (cases \bit k n\) case True have \signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\ by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True) then have \signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\ by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff) with True show ?thesis by (simp flip: minus_exp_eq_not_mask) next case False show ?thesis by (rule bit_eqI) (simp add: False bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq) qed lemma signed_take_bit_eq_take_bit_shift: \signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ for k :: int proof - have *: \take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\ by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff) have \take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\ by (simp add: minus_exp_eq_not_mask) also have \\ = take_bit n k OR NOT (mask n)\ by (rule disjunctive_add) (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff) finally have **: \take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\ . have \take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\ by (simp only: take_bit_add) also have \take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\ by (simp add: take_bit_Suc_from_most) finally have \take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\ by (simp add: ac_simps) also have \2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\ by (rule disjunctive_add) (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff) finally show ?thesis using * ** by (simp add: signed_take_bit_def concat_bit_Suc min_def ac_simps) qed lemma signed_take_bit_nonnegative_iff [simp]: \0 \ signed_take_bit n k \ \ bit k n\ for k :: int by (simp add: signed_take_bit_def not_less concat_bit_def) lemma signed_take_bit_negative_iff [simp]: \signed_take_bit n k < 0 \ bit k n\ for k :: int by (simp add: signed_take_bit_def not_less concat_bit_def) lemma signed_take_bit_int_eq_self_iff: \signed_take_bit n k = k \ - (2 ^ n) \ k \ k < 2 ^ n\ for k :: int by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps) +lemma signed_take_bit_int_eq_self: + \signed_take_bit n k = k\ if \- (2 ^ n) \ k\ \k < 2 ^ n\ + for k :: int + using that by (simp add: signed_take_bit_int_eq_self_iff) + lemma signed_take_bit_int_less_eq_self_iff: \signed_take_bit n k \ k \ - (2 ^ n) \ k\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_eq_self_iff algebra_simps) linarith lemma signed_take_bit_int_less_self_iff: \signed_take_bit n k < k \ 2 ^ n \ k\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_self_iff algebra_simps) lemma signed_take_bit_int_greater_self_iff: \k < signed_take_bit n k \ k < - (2 ^ n)\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_self_iff algebra_simps) linarith lemma signed_take_bit_int_greater_eq_self_iff: \k \ signed_take_bit n k \ k < 2 ^ n\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_eq_self_iff algebra_simps) lemma signed_take_bit_int_greater_eq: \k + 2 ^ Suc n \ signed_take_bit n k\ if \k < - (2 ^ n)\ for k :: int - using that take_bit_greater_eq [of \k + 2 ^ n\ \Suc n\] + using that take_bit_int_greater_eq [of \k + 2 ^ n\ \Suc n\] by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_int_less_eq: \signed_take_bit n k \ k - 2 ^ Suc n\ if \k \ 2 ^ n\ for k :: int - using that take_bit_less_eq [of \Suc n\ \k + 2 ^ n\] + using that take_bit_int_less_eq [of \Suc n\ \k + 2 ^ n\] by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_Suc_bit0 [simp]: \signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_bit1 [simp]: \signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_minus_bit0 [simp]: \signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_minus_bit1 [simp]: \signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_numeral_bit0 [simp]: \signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_bit1 [simp]: \signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_minus_bit0 [simp]: \signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_minus_bit1 [simp]: \signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_code [code]: \signed_take_bit n a = (let l = take_bit (Suc n) a in if bit l n then l + push_bit (Suc n) (- 1) else l)\ proof - have *: \take_bit (Suc n) a + push_bit n (- 2) = take_bit (Suc n) a OR NOT (mask (Suc n))\ by (auto simp add: bit_take_bit_iff bit_push_bit_iff bit_not_iff bit_mask_iff disjunctive_add simp flip: push_bit_minus_one_eq_not_mask) show ?thesis by (rule bit_eqI) (auto simp add: Let_def * bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq bit_not_iff bit_mask_iff bit_or_iff) qed subsection \Instance \<^typ>\nat\\ instantiation nat :: semiring_bit_operations begin definition and_nat :: \nat \ nat \ nat\ where \m AND n = nat (int m AND int n)\ for m n :: nat definition or_nat :: \nat \ nat \ nat\ where \m OR n = nat (int m OR int n)\ for m n :: nat definition xor_nat :: \nat \ nat \ nat\ where \m XOR n = nat (int m XOR int n)\ for m n :: nat definition mask_nat :: \nat \ nat\ where \mask n = (2 :: nat) ^ n - 1\ instance proof fix m n q :: nat show \bit (m AND n) q \ bit m q \ bit n q\ by (auto simp add: bit_nat_iff and_nat_def bit_and_iff less_le bit_eq_iff) show \bit (m OR n) q \ bit m q \ bit n q\ by (auto simp add: bit_nat_iff or_nat_def bit_or_iff less_le bit_eq_iff) show \bit (m XOR n) q \ bit m q \ bit n q\ by (auto simp add: bit_nat_iff xor_nat_def bit_xor_iff less_le bit_eq_iff) qed (simp add: mask_nat_def) end lemma and_nat_rec: \m AND n = of_bool (odd m \ odd n) + 2 * ((m div 2) AND (n div 2))\ for m n :: nat by (simp add: and_nat_def and_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma or_nat_rec: \m OR n = of_bool (odd m \ odd n) + 2 * ((m div 2) OR (n div 2))\ for m n :: nat by (simp add: or_nat_def or_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma xor_nat_rec: \m XOR n = of_bool (odd m \ odd n) + 2 * ((m div 2) XOR (n div 2))\ for m n :: nat by (simp add: xor_nat_def xor_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma Suc_0_and_eq [simp]: \Suc 0 AND n = n mod 2\ using one_and_eq [of n] by simp lemma and_Suc_0_eq [simp]: \n AND Suc 0 = n mod 2\ using and_one_eq [of n] by simp lemma Suc_0_or_eq: \Suc 0 OR n = n + of_bool (even n)\ using one_or_eq [of n] by simp lemma or_Suc_0_eq: \n OR Suc 0 = n + of_bool (even n)\ using or_one_eq [of n] by simp lemma Suc_0_xor_eq: \Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\ using one_xor_eq [of n] by simp lemma xor_Suc_0_eq: \n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\ using xor_one_eq [of n] by simp context semiring_bit_operations begin lemma of_nat_and_eq: \of_nat (m AND n) = of_nat m AND of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff) lemma of_nat_or_eq: \of_nat (m OR n) = of_nat m OR of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff) lemma of_nat_xor_eq: \of_nat (m XOR n) = of_nat m XOR of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff) end context ring_bit_operations begin lemma of_nat_mask_eq: \of_nat (mask n) = mask n\ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq) end subsection \Instances for \<^typ>\integer\ and \<^typ>\natural\\ unbundle integer.lifting natural.lifting instantiation integer :: ring_bit_operations begin lift_definition not_integer :: \integer \ integer\ is not . lift_definition and_integer :: \integer \ integer \ integer\ is \and\ . lift_definition or_integer :: \integer \ integer \ integer\ is or . lift_definition xor_integer :: \integer \ integer \ integer\ is xor . lift_definition mask_integer :: \nat \ integer\ is mask . instance by (standard; transfer) (simp_all add: minus_eq_not_minus_1 mask_eq_exp_minus_1 bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) end lemma [code]: \mask n = 2 ^ n - (1::integer)\ by (simp add: mask_eq_exp_minus_1) instantiation natural :: semiring_bit_operations begin lift_definition and_natural :: \natural \ natural \ natural\ is \and\ . lift_definition or_natural :: \natural \ natural \ natural\ is or . lift_definition xor_natural :: \natural \ natural \ natural\ is xor . lift_definition mask_natural :: \nat \ natural\ is mask . instance by (standard; transfer) (simp_all add: mask_eq_exp_minus_1 bit_and_iff bit_or_iff bit_xor_iff) end lemma [code]: \integer_of_natural (mask n) = mask n\ by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff) lifting_update integer.lifting lifting_forget integer.lifting lifting_update natural.lifting lifting_forget natural.lifting subsection \Key ideas of bit operations\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). \<^item> From this idea follows that \<^item> multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]} \<^item> Typical operations are characterized as follows: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]} \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]} \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]} \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]} \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm signed_take_bit_def [no_vars]} \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \ end diff --git a/src/HOL/Parity.thy b/src/HOL/Parity.thy --- a/src/HOL/Parity.thy +++ b/src/HOL/Parity.thy @@ -1,1925 +1,1947 @@ (* Title: HOL/Parity.thy Author: Jeremy Avigad Author: Jacques D. Fleuriot *) section \Parity in rings and semirings\ theory Parity imports Euclidean_Division begin subsection \Ring structures with parity and \even\/\odd\ predicates\ class semiring_parity = comm_semiring_1 + semiring_modulo + assumes even_iff_mod_2_eq_zero: "2 dvd a \ a mod 2 = 0" and odd_iff_mod_2_eq_one: "\ 2 dvd a \ a mod 2 = 1" and odd_one [simp]: "\ 2 dvd 1" begin abbreviation even :: "'a \ bool" where "even a \ 2 dvd a" abbreviation odd :: "'a \ bool" where "odd a \ \ 2 dvd a" lemma parity_cases [case_names even odd]: assumes "even a \ a mod 2 = 0 \ P" assumes "odd a \ a mod 2 = 1 \ P" shows P using assms by (cases "even a") (simp_all add: even_iff_mod_2_eq_zero [symmetric] odd_iff_mod_2_eq_one [symmetric]) lemma odd_of_bool_self [simp]: \odd (of_bool p) \ p\ by (cases p) simp_all lemma not_mod_2_eq_0_eq_1 [simp]: "a mod 2 \ 0 \ a mod 2 = 1" by (cases a rule: parity_cases) simp_all lemma not_mod_2_eq_1_eq_0 [simp]: "a mod 2 \ 1 \ a mod 2 = 0" by (cases a rule: parity_cases) simp_all lemma evenE [elim?]: assumes "even a" obtains b where "a = 2 * b" using assms by (rule dvdE) lemma oddE [elim?]: assumes "odd a" obtains b where "a = 2 * b + 1" proof - have "a = 2 * (a div 2) + a mod 2" by (simp add: mult_div_mod_eq) with assms have "a = 2 * (a div 2) + 1" by (simp add: odd_iff_mod_2_eq_one) then show ?thesis .. qed lemma mod_2_eq_odd: "a mod 2 = of_bool (odd a)" by (auto elim: oddE simp add: even_iff_mod_2_eq_zero) lemma of_bool_odd_eq_mod_2: "of_bool (odd a) = a mod 2" by (simp add: mod_2_eq_odd) lemma even_mod_2_iff [simp]: \even (a mod 2) \ even a\ by (simp add: mod_2_eq_odd) lemma mod2_eq_if: "a mod 2 = (if even a then 0 else 1)" by (simp add: mod_2_eq_odd) lemma even_zero [simp]: "even 0" by (fact dvd_0_right) lemma odd_even_add: "even (a + b)" if "odd a" and "odd b" proof - from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" by (blast elim: oddE) then have "a + b = 2 * c + 2 * d + (1 + 1)" by (simp only: ac_simps) also have "\ = 2 * (c + d + 1)" by (simp add: algebra_simps) finally show ?thesis .. qed lemma even_add [simp]: "even (a + b) \ (even a \ even b)" by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) lemma odd_add [simp]: "odd (a + b) \ \ (odd a \ odd b)" by simp lemma even_plus_one_iff [simp]: "even (a + 1) \ odd a" by (auto simp add: dvd_add_right_iff intro: odd_even_add) lemma even_mult_iff [simp]: "even (a * b) \ even a \ even b" (is "?P \ ?Q") proof assume ?Q then show ?P by auto next assume ?P show ?Q proof (rule ccontr) assume "\ (even a \ even b)" then have "odd a" and "odd b" by auto then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" by (blast elim: oddE) then have "a * b = (2 * r + 1) * (2 * s + 1)" by simp also have "\ = 2 * (2 * r * s + r + s) + 1" by (simp add: algebra_simps) finally have "odd (a * b)" by simp with \?P\ show False by auto qed qed lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" proof - have "even (2 * numeral n)" unfolding even_mult_iff by simp then have "even (numeral n + numeral n)" unfolding mult_2 . then show ?thesis unfolding numeral.simps . qed lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" proof assume "even (numeral (num.Bit1 n))" then have "even (numeral n + numeral n + 1)" unfolding numeral.simps . then have "even (2 * numeral n + 1)" unfolding mult_2 . then have "2 dvd numeral n * 2 + 1" by (simp add: ac_simps) then have "2 dvd 1" using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp then show False by simp qed lemma odd_numeral_BitM [simp]: \odd (numeral (Num.BitM w))\ by (cases w) simp_all lemma even_power [simp]: "even (a ^ n) \ even a \ n > 0" by (induct n) auto lemma mask_eq_sum_exp: \2 ^ n - 1 = (\m\{q. q < n}. 2 ^ m)\ proof - have *: \{q. q < Suc m} = insert m {q. q < m}\ for m by auto have \2 ^ n = (\m\{q. q < n}. 2 ^ m) + 1\ by (induction n) (simp_all add: ac_simps mult_2 *) then have \2 ^ n - 1 = (\m\{q. q < n}. 2 ^ m) + 1 - 1\ by simp then show ?thesis by simp qed end class ring_parity = ring + semiring_parity begin subclass comm_ring_1 .. lemma even_minus: "even (- a) \ even a" by (fact dvd_minus_iff) lemma even_diff [simp]: "even (a - b) \ even (a + b)" using even_add [of a "- b"] by simp end subsection \Special case: euclidean rings containing the natural numbers\ context unique_euclidean_semiring_with_nat begin subclass semiring_parity proof show "2 dvd a \ a mod 2 = 0" for a by (fact dvd_eq_mod_eq_0) show "\ 2 dvd a \ a mod 2 = 1" for a proof assume "a mod 2 = 1" then show "\ 2 dvd a" by auto next assume "\ 2 dvd a" have eucl: "euclidean_size (a mod 2) = 1" proof (rule order_antisym) show "euclidean_size (a mod 2) \ 1" using mod_size_less [of 2 a] by simp show "1 \ euclidean_size (a mod 2)" using \\ 2 dvd a\ by (simp add: Suc_le_eq dvd_eq_mod_eq_0) qed from \\ 2 dvd a\ have "\ of_nat 2 dvd division_segment a * of_nat (euclidean_size a)" by simp then have "\ of_nat 2 dvd of_nat (euclidean_size a)" by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment) then have "\ 2 dvd euclidean_size a" using of_nat_dvd_iff [of 2] by simp then have "euclidean_size a mod 2 = 1" by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0) then have "of_nat (euclidean_size a mod 2) = of_nat 1" by simp then have "of_nat (euclidean_size a) mod 2 = 1" by (simp add: of_nat_mod) from \\ 2 dvd a\ eucl show "a mod 2 = 1" by (auto intro: division_segment_eq_iff simp add: division_segment_mod) qed show "\ is_unit 2" proof (rule notI) assume "is_unit 2" then have "of_nat 2 dvd of_nat 1" by simp then have "is_unit (2::nat)" by (simp only: of_nat_dvd_iff) then show False by simp qed qed lemma even_of_nat [simp]: "even (of_nat a) \ even a" proof - have "even (of_nat a) \ of_nat 2 dvd of_nat a" by simp also have "\ \ even a" by (simp only: of_nat_dvd_iff) finally show ?thesis . qed lemma even_succ_div_two [simp]: "even a \ (a + 1) div 2 = a div 2" by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) lemma odd_succ_div_two [simp]: "odd a \ (a + 1) div 2 = a div 2 + 1" by (auto elim!: oddE simp add: add.assoc) lemma even_two_times_div_two: "even a \ 2 * (a div 2) = a" by (fact dvd_mult_div_cancel) lemma odd_two_times_div_two_succ [simp]: "odd a \ 2 * (a div 2) + 1 = a" using mult_div_mod_eq [of 2 a] by (simp add: even_iff_mod_2_eq_zero) lemma coprime_left_2_iff_odd [simp]: "coprime 2 a \ odd a" proof assume "odd a" show "coprime 2 a" proof (rule coprimeI) fix b assume "b dvd 2" "b dvd a" then have "b dvd a mod 2" by (auto intro: dvd_mod) with \odd a\ show "is_unit b" by (simp add: mod_2_eq_odd) qed next assume "coprime 2 a" show "odd a" proof (rule notI) assume "even a" then obtain b where "a = 2 * b" .. with \coprime 2 a\ have "coprime 2 (2 * b)" by simp moreover have "\ coprime 2 (2 * b)" by (rule not_coprimeI [of 2]) simp_all ultimately show False by blast qed qed lemma coprime_right_2_iff_odd [simp]: "coprime a 2 \ odd a" using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) end context unique_euclidean_ring_with_nat begin subclass ring_parity .. lemma minus_1_mod_2_eq [simp]: "- 1 mod 2 = 1" by (simp add: mod_2_eq_odd) lemma minus_1_div_2_eq [simp]: "- 1 div 2 = - 1" proof - from div_mult_mod_eq [of "- 1" 2] have "- 1 div 2 * 2 = - 1 * 2" using add_implies_diff by fastforce then show ?thesis using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp qed end subsection \Instance for \<^typ>\nat\\ instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) lemma even_Suc_Suc_iff [simp]: "even (Suc (Suc n)) \ even n" using dvd_add_triv_right_iff [of 2 n] by simp lemma even_Suc [simp]: "even (Suc n) \ odd n" using even_plus_one_iff [of n] by simp lemma even_diff_nat [simp]: "even (m - n) \ m < n \ even (m + n)" for m n :: nat proof (cases "n \ m") case True then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) moreover have "even (m - n) \ even (m - n + n * 2)" by simp ultimately have "even (m - n) \ even (m + n)" by (simp only:) then show ?thesis by auto next case False then show ?thesis by simp qed lemma odd_pos: "odd n \ 0 < n" for n :: nat by (auto elim: oddE) lemma Suc_double_not_eq_double: "Suc (2 * m) \ 2 * n" proof assume "Suc (2 * m) = 2 * n" moreover have "odd (Suc (2 * m))" and "even (2 * n)" by simp_all ultimately show False by simp qed lemma double_not_eq_Suc_double: "2 * m \ Suc (2 * n)" using Suc_double_not_eq_double [of n m] by simp lemma odd_Suc_minus_one [simp]: "odd n \ Suc (n - Suc 0) = n" by (auto elim: oddE) lemma even_Suc_div_two [simp]: "even n \ Suc n div 2 = n div 2" using even_succ_div_two [of n] by simp lemma odd_Suc_div_two [simp]: "odd n \ Suc n div 2 = Suc (n div 2)" using odd_succ_div_two [of n] by simp lemma odd_two_times_div_two_nat [simp]: assumes "odd n" shows "2 * (n div 2) = n - (1 :: nat)" proof - from assms have "2 * (n div 2) + 1 = n" by (rule odd_two_times_div_two_succ) then have "Suc (2 * (n div 2)) - 1 = n - 1" by simp then show ?thesis by simp qed lemma not_mod2_eq_Suc_0_eq_0 [simp]: "n mod 2 \ Suc 0 \ n mod 2 = 0" using not_mod_2_eq_1_eq_0 [of n] by simp lemma odd_card_imp_not_empty: \A \ {}\ if \odd (card A)\ using that by auto lemma nat_induct2 [case_names 0 1 step]: assumes "P 0" "P 1" and step: "\n::nat. P n \ P (n + 2)" shows "P n" proof (induct n rule: less_induct) case (less n) show ?case proof (cases "n < Suc (Suc 0)") case True then show ?thesis using assms by (auto simp: less_Suc_eq) next case False then obtain k where k: "n = Suc (Suc k)" by (force simp: not_less nat_le_iff_add) then have "k2 ^ n - Suc 0 = (\m\{q. q < n}. 2 ^ m)\ using mask_eq_sum_exp [where ?'a = nat] by simp context semiring_parity begin lemma even_sum_iff: \even (sum f A) \ even (card {a\A. odd (f a)})\ if \finite A\ using that proof (induction A) case empty then show ?case by simp next case (insert a A) moreover have \{b \ insert a A. odd (f b)} = (if odd (f a) then {a} else {}) \ {b \ A. odd (f b)}\ by auto ultimately show ?case by simp qed lemma even_prod_iff: \even (prod f A) \ (\a\A. even (f a))\ if \finite A\ using that by (induction A) simp_all lemma even_mask_iff [simp]: \even (2 ^ n - 1) \ n = 0\ proof (cases \n = 0\) case True then show ?thesis by simp next case False then have \{a. a = 0 \ a < n} = {0}\ by auto then show ?thesis by (auto simp add: mask_eq_sum_exp even_sum_iff) qed end subsection \Parity and powers\ context ring_1 begin lemma power_minus_even [simp]: "even n \ (- a) ^ n = a ^ n" by (auto elim: evenE) lemma power_minus_odd [simp]: "odd n \ (- a) ^ n = - (a ^ n)" by (auto elim: oddE) lemma uminus_power_if: "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" by auto lemma neg_one_even_power [simp]: "even n \ (- 1) ^ n = 1" by simp lemma neg_one_odd_power [simp]: "odd n \ (- 1) ^ n = - 1" by simp lemma neg_one_power_add_eq_neg_one_power_diff: "k \ n \ (- 1) ^ (n + k) = (- 1) ^ (n - k)" by (cases "even (n + k)") auto lemma minus_one_power_iff: "(- 1) ^ n = (if even n then 1 else - 1)" by (induct n) auto end context linordered_idom begin lemma zero_le_even_power: "even n \ 0 \ a ^ n" by (auto elim: evenE) lemma zero_le_odd_power: "odd n \ 0 \ a ^ n \ 0 \ a" by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) lemma zero_le_power_eq: "0 \ a ^ n \ even n \ odd n \ 0 \ a" by (auto simp add: zero_le_even_power zero_le_odd_power) lemma zero_less_power_eq: "0 < a ^ n \ n = 0 \ even n \ a \ 0 \ odd n \ 0 < a" proof - have [simp]: "0 = a ^ n \ a = 0 \ n > 0" unfolding power_eq_0_iff [of a n, symmetric] by blast show ?thesis unfolding less_le zero_le_power_eq by auto qed lemma power_less_zero_eq [simp]: "a ^ n < 0 \ odd n \ a < 0" unfolding not_le [symmetric] zero_le_power_eq by auto lemma power_le_zero_eq: "a ^ n \ 0 \ n > 0 \ (odd n \ a \ 0 \ even n \ a = 0)" unfolding not_less [symmetric] zero_less_power_eq by auto lemma power_even_abs: "even n \ \a\ ^ n = a ^ n" using power_abs [of a n] by (simp add: zero_le_even_power) lemma power_mono_even: assumes "even n" and "\a\ \ \b\" shows "a ^ n \ b ^ n" proof - have "0 \ \a\" by auto with \\a\ \ \b\\ have "\a\ ^ n \ \b\ ^ n" by (rule power_mono) with \even n\ show ?thesis by (simp add: power_even_abs) qed lemma power_mono_odd: assumes "odd n" and "a \ b" shows "a ^ n \ b ^ n" proof (cases "b < 0") case True with \a \ b\ have "- b \ - a" and "0 \ - b" by auto then have "(- b) ^ n \ (- a) ^ n" by (rule power_mono) with \odd n\ show ?thesis by simp next case False then have "0 \ b" by auto show ?thesis proof (cases "a < 0") case True then have "n \ 0" and "a \ 0" using \odd n\ [THEN odd_pos] by auto then have "a ^ n \ 0" unfolding power_le_zero_eq using \odd n\ by auto moreover from \0 \ b\ have "0 \ b ^ n" by auto ultimately show ?thesis by auto next case False then have "0 \ a" by auto with \a \ b\ show ?thesis using power_mono by auto qed qed text \Simplify, when the exponent is a numeral\ lemma zero_le_power_eq_numeral [simp]: "0 \ a ^ numeral w \ even (numeral w :: nat) \ odd (numeral w :: nat) \ 0 \ a" by (fact zero_le_power_eq) lemma zero_less_power_eq_numeral [simp]: "0 < a ^ numeral w \ numeral w = (0 :: nat) \ even (numeral w :: nat) \ a \ 0 \ odd (numeral w :: nat) \ 0 < a" by (fact zero_less_power_eq) lemma power_le_zero_eq_numeral [simp]: "a ^ numeral w \ 0 \ (0 :: nat) < numeral w \ (odd (numeral w :: nat) \ a \ 0 \ even (numeral w :: nat) \ a = 0)" by (fact power_le_zero_eq) lemma power_less_zero_eq_numeral [simp]: "a ^ numeral w < 0 \ odd (numeral w :: nat) \ a < 0" by (fact power_less_zero_eq) lemma power_even_abs_numeral [simp]: "even (numeral w :: nat) \ \a\ ^ numeral w = a ^ numeral w" by (fact power_even_abs) end context unique_euclidean_semiring_with_nat begin lemma even_mask_div_iff': \even ((2 ^ m - 1) div 2 ^ n) \ m \ n\ proof - have \even ((2 ^ m - 1) div 2 ^ n) \ even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\ by (simp only: of_nat_div) (simp add: of_nat_diff) also have \\ \ even ((2 ^ m - Suc 0) div 2 ^ n)\ by simp also have \\ \ m \ n\ proof (cases \m \ n\) case True then show ?thesis by (simp add: Suc_le_lessD) next case False then obtain r where r: \m = n + Suc r\ using less_imp_Suc_add by fastforce from r have \{q. q < m} \ {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \ q \ q < m}\ by (auto simp add: dvd_power_iff_le) moreover from r have \{q. q < m} \ {q. \ 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\ by (auto simp add: dvd_power_iff_le) moreover from False have \{q. n \ q \ q < m \ q \ n} = {n}\ by auto then have \odd ((\a\{q. n \ q \ q < m}. 2 ^ a div (2::nat) ^ n) + sum ((^) 2) {q. q < n} div 2 ^ n)\ by (simp_all add: euclidean_semiring_cancel_class.power_diff_power_eq semiring_parity_class.even_sum_iff not_less mask_eq_sum_exp_nat [symmetric]) ultimately have \odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\ by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all with False show ?thesis by (simp add: mask_eq_sum_exp_nat) qed finally show ?thesis . qed end subsection \Instance for \<^typ>\int\\ lemma even_diff_iff: "even (k - l) \ even (k + l)" for k l :: int by (fact even_diff) lemma even_abs_add_iff: "even (\k\ + l) \ even (k + l)" for k l :: int by simp lemma even_add_abs_iff: "even (k + \l\) \ even (k + l)" for k l :: int by simp lemma even_nat_iff: "0 \ k \ even (nat k) \ even k" by (simp add: even_of_nat [of "nat k", where ?'a = int, symmetric]) 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 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 context assumes "SORT_CONSTRAINT('a::division_ring)" begin lemma power_int_minus_left: "power_int (-a :: 'a) n = (if even n then power_int a n else -power_int a n)" by (auto simp: power_int_def minus_one_power_iff even_nat_iff) lemma power_int_minus_left_even [simp]: "even n \ power_int (-a :: 'a) n = power_int a n" by (simp add: power_int_minus_left) lemma power_int_minus_left_odd [simp]: "odd n \ power_int (-a :: 'a) n = -power_int a n" by (simp add: power_int_minus_left) lemma power_int_minus_left_distrib: "NO_MATCH (-1) x \ power_int (-a :: 'a) n = power_int (-1) n * power_int a n" by (simp add: power_int_minus_left) lemma power_int_minus_one_minus: "power_int (-1 :: 'a) (-n) = power_int (-1) n" by (simp add: power_int_minus_left) lemma power_int_minus_one_diff_commute: "power_int (-1 :: 'a) (a - b) = power_int (-1) (b - a)" by (subst power_int_minus_one_minus [symmetric]) auto lemma power_int_minus_one_mult_self [simp]: "power_int (-1 :: 'a) m * power_int (-1) m = 1" by (simp add: power_int_minus_left) lemma power_int_minus_one_mult_self' [simp]: "power_int (-1 :: 'a) m * (power_int (-1) m * b) = b" by (simp add: power_int_minus_left) end subsection \Abstract bit structures\ class semiring_bits = semiring_parity + assumes bits_induct [case_names stable rec]: \(\a. a div 2 = a \ P a) \ (\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)) \ P a\ assumes bits_div_0 [simp]: \0 div a = 0\ and bits_div_by_1 [simp]: \a div 1 = a\ and bits_mod_div_trivial [simp]: \a mod b div b = 0\ and even_succ_div_2 [simp]: \even a \ (1 + a) div 2 = a div 2\ and even_mask_div_iff: \even ((2 ^ m - 1) div 2 ^ n) \ 2 ^ n = 0 \ m \ n\ and exp_div_exp_eq: \2 ^ m div 2 ^ n = of_bool (2 ^ m \ 0 \ m \ n) * 2 ^ (m - n)\ and div_exp_eq: \a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\ and mod_exp_eq: \a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\ and mult_exp_mod_exp_eq: \m \ n \ (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\ and div_exp_mod_exp_eq: \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ and even_mult_exp_div_exp_iff: \even (a * 2 ^ m div 2 ^ n) \ m > n \ 2 ^ n = 0 \ (m \ n \ even (a div 2 ^ (n - m)))\ fixes bit :: \'a \ nat \ bool\ assumes bit_iff_odd: \bit a n \ odd (a div 2 ^ n)\ begin text \ Having \<^const>\bit\ as definitional class operation takes into account that specific instances can be implemented differently wrt. code generation. \ lemma bits_div_by_0 [simp]: \a div 0 = 0\ by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero) lemma bits_1_div_2 [simp]: \1 div 2 = 0\ using even_succ_div_2 [of 0] by simp lemma bits_1_div_exp [simp]: \1 div 2 ^ n = of_bool (n = 0)\ using div_exp_eq [of 1 1] by (cases n) simp_all lemma even_succ_div_exp [simp]: \(1 + a) div 2 ^ n = a div 2 ^ n\ if \even a\ and \n > 0\ proof (cases n) case 0 with that show ?thesis by simp next case (Suc n) with \even a\ have \(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\ proof (induction n) case 0 then show ?case by simp next case (Suc n) then show ?case using div_exp_eq [of _ 1 \Suc n\, symmetric] by simp qed with Suc show ?thesis by simp qed lemma even_succ_mod_exp [simp]: \(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\ if \even a\ and \n > 0\ using div_mult_mod_eq [of \1 + a\ \2 ^ n\] that apply simp by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq) lemma bits_mod_by_1 [simp]: \a mod 1 = 0\ using div_mult_mod_eq [of a 1] by simp lemma bits_mod_0 [simp]: \0 mod a = 0\ using div_mult_mod_eq [of 0 a] by simp lemma bits_one_mod_two_eq_one [simp]: \1 mod 2 = 1\ by (simp add: mod2_eq_if) lemma bit_0 [simp]: \bit a 0 \ odd a\ by (simp add: bit_iff_odd) lemma bit_Suc: \bit a (Suc n) \ bit (a div 2) n\ using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd) lemma bit_rec: \bit a n \ (if n = 0 then odd a else bit (a div 2) (n - 1))\ by (cases n) (simp_all add: bit_Suc) lemma bit_0_eq [simp]: \bit 0 = bot\ by (simp add: fun_eq_iff bit_iff_odd) context fixes a assumes stable: \a div 2 = a\ begin lemma bits_stable_imp_add_self: \a + a mod 2 = 0\ proof - have \a div 2 * 2 + a mod 2 = a\ by (fact div_mult_mod_eq) then have \a * 2 + a mod 2 = a\ by (simp add: stable) then show ?thesis by (simp add: mult_2_right ac_simps) qed lemma stable_imp_bit_iff_odd: \bit a n \ odd a\ by (induction n) (simp_all add: stable bit_Suc) end lemma bit_iff_idd_imp_stable: \a div 2 = a\ if \\n. bit a n \ odd a\ using that proof (induction a rule: bits_induct) case (stable a) then show ?case by simp next case (rec a b) from rec.prems [of 1] have [simp]: \b = odd a\ by (simp add: rec.hyps bit_Suc) from rec.hyps have hyp: \(of_bool (odd a) + 2 * a) div 2 = a\ by simp have \bit a n \ odd a\ for n using rec.prems [of \Suc n\] by (simp add: hyp bit_Suc) then have \a div 2 = a\ by (rule rec.IH) then have \of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\ by (simp add: ac_simps) also have \\ = a\ using mult_div_mod_eq [of 2 a] by (simp add: of_bool_odd_eq_mod_2) finally show ?case using \a div 2 = a\ by (simp add: hyp) qed lemma exp_eq_0_imp_not_bit: \\ bit a n\ if \2 ^ n = 0\ using that by (simp add: bit_iff_odd) lemma bit_eqI: \a = b\ if \\n. 2 ^ n \ 0 \ bit a n \ bit b n\ proof - have \bit a n \ bit b n\ for n proof (cases \2 ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False then show ?thesis by (rule that) qed then show ?thesis proof (induction a arbitrary: b rule: bits_induct) case (stable a) from stable(2) [of 0] have **: \even b \ even a\ by simp have \b div 2 = b\ proof (rule bit_iff_idd_imp_stable) fix n from stable have *: \bit b n \ bit a n\ by simp also have \bit a n \ odd a\ using stable by (simp add: stable_imp_bit_iff_odd) finally show \bit b n \ odd b\ by (simp add: **) qed from ** have \a mod 2 = b mod 2\ by (simp add: mod2_eq_if) then have \a mod 2 + (a + b) = b mod 2 + (a + b)\ by simp then have \a + a mod 2 + b = b + b mod 2 + a\ by (simp add: ac_simps) with \a div 2 = a\ \b div 2 = b\ show ?case by (simp add: bits_stable_imp_add_self) next case (rec a p) from rec.prems [of 0] have [simp]: \p = odd b\ by simp from rec.hyps have \bit a n \ bit (b div 2) n\ for n using rec.prems [of \Suc n\] by (simp add: bit_Suc) then have \a = b div 2\ by (rule rec.IH) then have \2 * a = 2 * (b div 2)\ by simp then have \b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\ by simp also have \\ = b\ by (fact mod_mult_div_eq) finally show ?case by (auto simp add: mod2_eq_if) qed qed lemma bit_eq_iff: \a = b \ (\n. bit a n \ bit b n)\ by (auto intro: bit_eqI) lemma bit_exp_iff: \bit (2 ^ m) n \ 2 ^ m \ 0 \ m = n\ by (auto simp add: bit_iff_odd exp_div_exp_eq) lemma bit_1_iff: \bit 1 n \ 1 \ 0 \ n = 0\ using bit_exp_iff [of 0 n] by simp lemma bit_2_iff: \bit 2 n \ 2 \ 0 \ n = 1\ using bit_exp_iff [of 1 n] by auto lemma even_bit_succ_iff: \bit (1 + a) n \ bit a n \ n = 0\ if \even a\ using that by (cases \n = 0\) (simp_all add: bit_iff_odd) lemma odd_bit_iff_bit_pred: \bit a n \ bit (a - 1) n \ n = 0\ if \odd a\ proof - from \odd a\ obtain b where \a = 2 * b + 1\ .. moreover have \bit (2 * b) n \ n = 0 \ bit (1 + 2 * b) n\ using even_bit_succ_iff by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma bit_double_iff: \bit (2 * a) n \ bit a (n - 1) \ n \ 0 \ 2 ^ n \ 0\ using even_mult_exp_div_exp_iff [of a 1 n] by (cases n, auto simp add: bit_iff_odd ac_simps) lemma bit_eq_rec: \a = b \ (even a \ even b) \ a div 2 = b div 2\ (is \?P = ?Q\) proof assume ?P then show ?Q by simp next assume ?Q then have \even a \ even b\ and \a div 2 = b div 2\ by simp_all show ?P proof (rule bit_eqI) fix n show \bit a n \ bit b n\ proof (cases n) case 0 with \even a \ even b\ show ?thesis by simp next case (Suc n) moreover from \a div 2 = b div 2\ have \bit (a div 2) n = bit (b div 2) n\ by simp ultimately show ?thesis by (simp add: bit_Suc) qed qed qed lemma bit_mod_2_iff [simp]: \bit (a mod 2) n \ n = 0 \ odd a\ by (cases a rule: parity_cases) (simp_all add: bit_iff_odd) lemma bit_mask_iff: \bit (2 ^ m - 1) n \ 2 ^ n \ 0 \ n < m\ by (simp add: bit_iff_odd even_mask_div_iff not_le) lemma bit_Numeral1_iff [simp]: \bit (numeral Num.One) n \ n = 0\ by (simp add: bit_rec) lemma exp_add_not_zero_imp: \2 ^ m \ 0\ and \2 ^ n \ 0\ if \2 ^ (m + n) \ 0\ proof - have \\ (2 ^ m = 0 \ 2 ^ n = 0)\ proof (rule notI) assume \2 ^ m = 0 \ 2 ^ n = 0\ then have \2 ^ (m + n) = 0\ by (rule disjE) (simp_all add: power_add) with that show False .. qed then show \2 ^ m \ 0\ and \2 ^ n \ 0\ by simp_all qed lemma bit_disjunctive_add_iff: \bit (a + b) n \ bit a n \ bit b n\ if \\n. \ bit a n \ \ bit b n\ proof (cases \2 ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False with that show ?thesis proof (induction n arbitrary: a b) case 0 from "0.prems"(1) [of 0] show ?case by auto next case (Suc n) from Suc.prems(1) [of 0] have even: \even a \ even b\ by auto have bit: \\ bit (a div 2) n \ \ bit (b div 2) n\ for n using Suc.prems(1) [of \Suc n\] by (simp add: bit_Suc) from Suc.prems(2) have \2 * 2 ^ n \ 0\ \2 ^ n \ 0\ by (auto simp add: mult_2) have \a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\ using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp also have \\ = of_bool (odd a \ odd b) + 2 * (a div 2 + b div 2)\ using even by (auto simp add: algebra_simps mod2_eq_if) finally have \bit ((a + b) div 2) n \ bit (a div 2 + b div 2) n\ using \2 * 2 ^ n \ 0\ by simp (simp flip: bit_Suc add: bit_double_iff) also have \\ \ bit (a div 2) n \ bit (b div 2) n\ using bit \2 ^ n \ 0\ by (rule Suc.IH) finally show ?case by (simp add: bit_Suc) qed qed lemma exp_add_not_zero_imp_left: \2 ^ m \ 0\ and exp_add_not_zero_imp_right: \2 ^ n \ 0\ if \2 ^ (m + n) \ 0\ proof - have \\ (2 ^ m = 0 \ 2 ^ n = 0)\ proof (rule notI) assume \2 ^ m = 0 \ 2 ^ n = 0\ then have \2 ^ (m + n) = 0\ by (rule disjE) (simp_all add: power_add) with that show False .. qed then show \2 ^ m \ 0\ and \2 ^ n \ 0\ by simp_all qed lemma exp_not_zero_imp_exp_diff_not_zero: \2 ^ (n - m) \ 0\ if \2 ^ n \ 0\ proof (cases \m \ n\) case True moreover define q where \q = n - m\ ultimately have \n = m + q\ by simp with that show ?thesis by (simp add: exp_add_not_zero_imp_right) next case False with that show ?thesis by simp qed end lemma nat_bit_induct [case_names zero even odd]: "P n" if zero: "P 0" and even: "\n. P n \ n > 0 \ P (2 * n)" and odd: "\n. P n \ P (Suc (2 * n))" proof (induction n rule: less_induct) case (less n) show "P n" proof (cases "n = 0") case True with zero show ?thesis by simp next case False with less have hyp: "P (n div 2)" by simp show ?thesis proof (cases "even n") case True then have "n \ 1" by auto with \n \ 0\ have "n div 2 > 0" by simp with \even n\ hyp even [of "n div 2"] show ?thesis by simp next case False with hyp odd [of "n div 2"] show ?thesis by simp qed qed qed instantiation nat :: semiring_bits begin definition bit_nat :: \nat \ nat \ bool\ where \bit_nat m n \ odd (m div 2 ^ n)\ instance proof show \P n\ if stable: \\n. n div 2 = n \ P n\ and rec: \\n b. P n \ (of_bool b + 2 * n) div 2 = n \ P (of_bool b + 2 * n)\ for P and n :: nat proof (induction n rule: nat_bit_induct) case zero from stable [of 0] show ?case by simp next case (even n) with rec [of n False] show ?case by simp next case (odd n) with rec [of n True] show ?case by simp qed show \q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\ for q m n :: nat apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin) apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes) done show \(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\ if \m \ n\ for q m n :: nat using that apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin) apply (simp add: mult.commute) done show \even ((2 ^ m - (1::nat)) div 2 ^ n) \ 2 ^ n = (0::nat) \ m \ n\ for m n :: nat using even_mask_div_iff' [where ?'a = nat, of m n] by simp show \even (q * 2 ^ m div 2 ^ n) \ n < m \ (2::nat) ^ n = 0 \ m \ n \ even (q div 2 ^ (n - m))\ for m n q r :: nat apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc) done qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff bit_nat_def) end lemma int_bit_induct [case_names zero minus even odd]: "P k" if zero_int: "P 0" and minus_int: "P (- 1)" and even_int: "\k. P k \ k \ 0 \ P (k * 2)" and odd_int: "\k. P k \ k \ - 1 \ P (1 + (k * 2))" for k :: int proof (cases "k \ 0") case True define n where "n = nat k" with True have "k = int n" by simp then show "P k" proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: zero_int) next case (even n) have "P (int n * 2)" by (rule even_int) (use even in simp_all) with even show ?case by (simp add: ac_simps) next case (odd n) have "P (1 + (int n * 2))" by (rule odd_int) (use odd in simp_all) with odd show ?case by (simp add: ac_simps) qed next case False define n where "n = nat (- k - 1)" with False have "k = - int n - 1" by simp then show "P k" proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: minus_int) next case (even n) have "P (1 + (- int (Suc n) * 2))" by (rule odd_int) (use even in \simp_all add: algebra_simps\) also have "\ = - int (2 * n) - 1" by (simp add: algebra_simps) finally show ?case using even.prems by simp next case (odd n) have "P (- int (Suc n) * 2)" by (rule even_int) (use odd in \simp_all add: algebra_simps\) also have "\ = - int (Suc (2 * n)) - 1" by (simp add: algebra_simps) finally show ?case using odd.prems by simp qed qed context semiring_bits begin lemma even_of_nat_iff: \even (of_nat n) \ even n\ by (induction n rule: nat_bit_induct) simp_all lemma bit_of_nat_iff: \bit (of_nat m) n \ (2::'a) ^ n \ 0 \ bit m n\ proof (cases \(2::'a) ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False then have \bit (of_nat m) n \ bit m n\ proof (induction m arbitrary: n rule: nat_bit_induct) case zero then show ?case by simp next case (even m) then show ?case by (cases n) (auto simp add: bit_double_iff Parity.bit_double_iff dest: mult_not_zero) next case (odd m) then show ?case by (cases n) (auto simp add: bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero) qed with False show ?thesis by simp qed end instantiation int :: semiring_bits begin definition bit_int :: \int \ nat \ bool\ where \bit_int k n \ odd (k div 2 ^ n)\ instance proof show \P k\ if stable: \\k. k div 2 = k \ P k\ and rec: \\k b. P k \ (of_bool b + 2 * k) div 2 = k \ P (of_bool b + 2 * k)\ for P and k :: int proof (induction k rule: int_bit_induct) case zero from stable [of 0] show ?case by simp next case minus from stable [of \- 1\] show ?case by simp next case (even k) with rec [of k False] show ?case by (simp add: ac_simps) next case (odd k) with rec [of k True] show ?case by (simp add: ac_simps) qed show \(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat proof (cases \m < n\) case True then have \n = m + (n - m)\ by simp then have \(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\ by simp also have \\ = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\ by (simp add: power_add) also have \\ = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\ by (simp add: zdiv_zmult2_eq) finally show ?thesis using \m < n\ by simp next case False then show ?thesis by (simp add: power_diff) qed show \k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\ for m n :: nat and k :: int using mod_exp_eq [of \nat k\ m n] apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin) apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add) apply (simp only: flip: mult.left_commute [of \2 ^ m\]) apply (subst zmod_zmult2_eq) apply simp_all done show \(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\ if \m \ n\ for m n :: nat and k :: int using that apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin) apply (simp add: ac_simps) done show \even ((2 ^ m - (1::int)) div 2 ^ n) \ 2 ^ n = (0::int) \ m \ n\ for m n :: nat using even_mask_div_iff' [where ?'a = int, of m n] by simp show \even (k * 2 ^ m div 2 ^ n) \ n < m \ (2::int) ^ n = 0 \ m \ n \ even (k div 2 ^ (n - m))\ for m n :: nat and k l :: int apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2)) done qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le bit_int_def) end class semiring_bit_shifts = semiring_bits + fixes push_bit :: \nat \ 'a \ 'a\ assumes push_bit_eq_mult: \push_bit n a = a * 2 ^ n\ fixes drop_bit :: \nat \ 'a \ 'a\ assumes drop_bit_eq_div: \drop_bit n a = a div 2 ^ n\ fixes take_bit :: \nat \ 'a \ 'a\ assumes take_bit_eq_mod: \take_bit n a = a mod 2 ^ n\ begin text \ Logically, \<^const>\push_bit\, \<^const>\drop_bit\ and \<^const>\take_bit\ are just aliases; having them as separate operations makes proofs easier, otherwise proof automation would fiddle with concrete expressions \<^term>\2 ^ n\ in a way obfuscating the basic algebraic relationships between those operations. Having them as definitional class operations takes into account that specific instances of these can be implemented differently wrt. code generation. \ lemma bit_iff_odd_drop_bit: \bit a n \ odd (drop_bit n a)\ by (simp add: bit_iff_odd drop_bit_eq_div) lemma even_drop_bit_iff_not_bit: \even (drop_bit n a) \ \ bit a n\ by (simp add: bit_iff_odd_drop_bit) lemma div_push_bit_of_1_eq_drop_bit: \a div push_bit n 1 = drop_bit n a\ by (simp add: push_bit_eq_mult drop_bit_eq_div) lemma bits_ident: "push_bit n (drop_bit n a) + take_bit n a = a" using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) lemma push_bit_push_bit [simp]: "push_bit m (push_bit n a) = push_bit (m + n) a" by (simp add: push_bit_eq_mult power_add ac_simps) lemma push_bit_0_id [simp]: "push_bit 0 = id" by (simp add: fun_eq_iff push_bit_eq_mult) lemma push_bit_of_0 [simp]: "push_bit n 0 = 0" by (simp add: push_bit_eq_mult) lemma push_bit_of_1: "push_bit n 1 = 2 ^ n" by (simp add: push_bit_eq_mult) lemma push_bit_Suc [simp]: "push_bit (Suc n) a = push_bit n (a * 2)" by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_double: "push_bit n (a * 2) = push_bit n a * 2" by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_add: "push_bit n (a + b) = push_bit n a + push_bit n b" by (simp add: push_bit_eq_mult algebra_simps) lemma push_bit_numeral [simp]: \push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\ by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) lemma take_bit_0 [simp]: "take_bit 0 a = 0" by (simp add: take_bit_eq_mod) lemma take_bit_Suc: \take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\ proof - have \take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\ using even_succ_mod_exp [of \2 * (a div 2)\ \Suc n\] mult_exp_mod_exp_eq [of 1 \Suc n\ \a div 2\] by (auto simp add: take_bit_eq_mod ac_simps) then show ?thesis using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd) qed lemma take_bit_rec: \take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\ by (cases n) (simp_all add: take_bit_Suc) lemma take_bit_Suc_0 [simp]: \take_bit (Suc 0) a = a mod 2\ by (simp add: take_bit_eq_mod) lemma take_bit_of_0 [simp]: "take_bit n 0 = 0" by (simp add: take_bit_eq_mod) lemma take_bit_of_1 [simp]: "take_bit n 1 = of_bool (n > 0)" by (cases n) (simp_all add: take_bit_Suc) lemma drop_bit_of_0 [simp]: "drop_bit n 0 = 0" by (simp add: drop_bit_eq_div) lemma drop_bit_of_1 [simp]: "drop_bit n 1 = of_bool (n = 0)" by (simp add: drop_bit_eq_div) lemma drop_bit_0 [simp]: "drop_bit 0 = id" by (simp add: fun_eq_iff drop_bit_eq_div) lemma drop_bit_Suc: "drop_bit (Suc n) a = drop_bit n (a div 2)" using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div) lemma drop_bit_rec: "drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))" by (cases n) (simp_all add: drop_bit_Suc) lemma drop_bit_half: "drop_bit n (a div 2) = drop_bit n a div 2" by (induction n arbitrary: a) (simp_all add: drop_bit_Suc) lemma drop_bit_of_bool [simp]: "drop_bit n (of_bool b) = of_bool (n = 0 \ b)" by (cases n) simp_all lemma even_take_bit_eq [simp]: \even (take_bit n a) \ n = 0 \ even a\ by (simp add: take_bit_rec [of n a]) lemma take_bit_take_bit [simp]: "take_bit m (take_bit n a) = take_bit (min m n) a" by (simp add: take_bit_eq_mod mod_exp_eq ac_simps) lemma drop_bit_drop_bit [simp]: "drop_bit m (drop_bit n a) = drop_bit (m + n) a" by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps) lemma push_bit_take_bit: "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps) using mult_exp_mod_exp_eq [of m \m + n\ a] apply (simp add: ac_simps power_add) done lemma take_bit_push_bit: "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" proof (cases "m \ n") case True then show ?thesis apply (simp add:) apply (simp_all add: push_bit_eq_mult take_bit_eq_mod) apply (auto dest!: le_Suc_ex simp add: power_add ac_simps) using mult_exp_mod_exp_eq [of m m \a * 2 ^ n\ for n] apply (simp add: ac_simps) done next case False then show ?thesis using push_bit_take_bit [of n "m - n" a] by simp qed lemma take_bit_drop_bit: "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq) lemma drop_bit_take_bit: "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" proof (cases "m \ n") case True then show ?thesis using take_bit_drop_bit [of "n - m" m a] by simp next case False then obtain q where \m = n + q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \drop_bit m (take_bit n a) = 0\ using div_exp_eq [of \a mod 2 ^ n\ n q] by (simp add: take_bit_eq_mod drop_bit_eq_div) with False show ?thesis by simp qed lemma even_push_bit_iff [simp]: \even (push_bit n a) \ n \ 0 \ even a\ by (simp add: push_bit_eq_mult) auto lemma bit_push_bit_iff: \bit (push_bit m a) n \ m \ n \ 2 ^ n \ 0 \ bit a (n - m)\ by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff) lemma bit_drop_bit_eq: \bit (drop_bit n a) = bit a \ (+) n\ by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div) lemma bit_take_bit_iff: \bit (take_bit m a) n \ n < m \ bit a n\ by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div) lemma stable_imp_drop_bit_eq: \drop_bit n a = a\ if \a div 2 = a\ by (induction n) (simp_all add: that drop_bit_Suc) lemma stable_imp_take_bit_eq: \take_bit n a = (if even a then 0 else 2 ^ n - 1)\ if \a div 2 = a\ proof (rule bit_eqI) fix m assume \2 ^ m \ 0\ with that show \bit (take_bit n a) m \ bit (if even a then 0 else 2 ^ n - 1) m\ by (simp add: bit_take_bit_iff bit_mask_iff stable_imp_bit_iff_odd) qed lemma exp_dvdE: assumes \2 ^ n dvd a\ obtains b where \a = push_bit n b\ proof - from assms obtain b where \a = 2 ^ n * b\ .. then have \a = push_bit n b\ by (simp add: push_bit_eq_mult ac_simps) with that show thesis . qed lemma take_bit_eq_0_iff: \take_bit n a = 0 \ 2 ^ n dvd a\ (is \?P \ ?Q\) proof assume ?P then show ?Q by (simp add: take_bit_eq_mod mod_0_imp_dvd) next assume ?Q then obtain b where \a = push_bit n b\ by (rule exp_dvdE) then show ?P by (simp add: take_bit_push_bit) qed lemma take_bit_tightened: \take_bit m a = take_bit m b\ if \take_bit n a = take_bit n b\ and \m \ n\ proof - from that have \take_bit m (take_bit n a) = take_bit m (take_bit n b)\ by simp then have \take_bit (min m n) a = take_bit (min m n) b\ by simp with that show ?thesis by (simp add: min_def) qed end instantiation nat :: semiring_bit_shifts begin definition push_bit_nat :: \nat \ nat \ nat\ where \push_bit_nat n m = m * 2 ^ n\ definition drop_bit_nat :: \nat \ nat \ nat\ where \drop_bit_nat n m = m div 2 ^ n\ definition take_bit_nat :: \nat \ nat \ nat\ where \take_bit_nat n m = m mod 2 ^ n\ instance by standard (simp_all add: push_bit_nat_def drop_bit_nat_def take_bit_nat_def) end context semiring_bit_shifts begin lemma push_bit_of_nat: \push_bit n (of_nat m) = of_nat (push_bit n m)\ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) lemma of_nat_push_bit: \of_nat (push_bit m n) = push_bit m (of_nat n)\ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) lemma take_bit_of_nat: \take_bit n (of_nat m) = of_nat (take_bit n m)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff) lemma of_nat_take_bit: \of_nat (take_bit n m) = take_bit n (of_nat m)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff) end instantiation int :: semiring_bit_shifts begin definition push_bit_int :: \nat \ int \ int\ where \push_bit_int n k = k * 2 ^ n\ definition drop_bit_int :: \nat \ int \ int\ where \drop_bit_int n k = k div 2 ^ n\ definition take_bit_int :: \nat \ int \ int\ where \take_bit_int n k = k mod 2 ^ n\ instance by standard (simp_all add: push_bit_int_def drop_bit_int_def take_bit_int_def) end lemma bit_push_bit_iff_nat: \bit (push_bit m q) n \ m \ n \ bit q (n - m)\ for q :: nat by (auto simp add: bit_push_bit_iff) lemma bit_push_bit_iff_int: \bit (push_bit m k) n \ m \ n \ bit k (n - m)\ for k :: int by (auto simp add: bit_push_bit_iff) -lemma take_bit_nat_less_exp: +lemma take_bit_nat_less_exp [simp]: \take_bit n m < 2 ^ n\ for n m ::nat by (simp add: take_bit_eq_mod) +lemma take_bit_nonnegative [simp]: + \take_bit n k \ 0\ for k :: int + by (simp add: take_bit_eq_mod) + +lemma not_take_bit_negative [simp]: + \\ take_bit n k < 0\ for k :: int + by (simp add: not_less) + +lemma take_bit_int_less_exp [simp]: + \take_bit n k < 2 ^ n\ for k :: int + by (simp add: take_bit_eq_mod) + lemma take_bit_nat_eq_self_iff: \take_bit n m = m \ m < 2 ^ n\ (is \?P \ ?Q\) for n m :: nat proof assume ?P moreover note take_bit_nat_less_exp [of n m] ultimately show ?Q by simp next assume ?Q then show ?P by (simp add: take_bit_eq_mod) qed -lemma take_bit_nat_less_eq_self: - \take_bit n m \ m\ for n m :: nat - by (simp add: take_bit_eq_mod) - -lemma take_bit_nonnegative [simp]: - \take_bit n k \ 0\ - for k :: int - by (simp add: take_bit_eq_mod) - -lemma not_take_bit_negative [simp]: - \\ take_bit n k < 0\ - for k :: int - by (simp add: not_less) - -lemma take_bit_int_less_exp: - \take_bit n k < 2 ^ n\ for k :: int - by (simp add: take_bit_eq_mod) +lemma take_bit_nat_eq_self: + \take_bit n m = m\ if \m < 2 ^ n\ for m n :: nat + using that by (simp add: take_bit_nat_eq_self_iff) lemma take_bit_int_eq_self_iff: \take_bit n k = k \ 0 \ k \ k < 2 ^ n\ (is \?P \ ?Q\) for k :: int proof assume ?P moreover note take_bit_int_less_exp [of n k] take_bit_nonnegative [of n k] ultimately show ?Q by simp next assume ?Q then show ?P by (simp add: take_bit_eq_mod) qed +lemma take_bit_int_eq_self: + \take_bit n k = k\ if \0 \ k\ \k < 2 ^ n\ for k :: int + using that by (simp add: take_bit_int_eq_self_iff) + +lemma take_bit_nat_less_eq_self [simp]: + \take_bit n m \ m\ for n m :: nat + by (simp add: take_bit_eq_mod) + +lemma take_bit_nat_less_self_iff: + \take_bit n m < m \ 2 ^ n \ m\ (is \?P \ ?Q\) + for m n :: nat +proof + assume ?P + then have \take_bit n m \ m\ + by simp + then show \?Q\ + by (simp add: take_bit_nat_eq_self_iff) +next + have \take_bit n m < 2 ^ n\ + by (fact take_bit_nat_less_exp) + also assume ?Q + finally show ?P . +qed + class unique_euclidean_semiring_with_bit_shifts = unique_euclidean_semiring_with_nat + semiring_bit_shifts begin lemma take_bit_of_exp [simp]: \take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\ by (simp add: take_bit_eq_mod exp_mod_exp) lemma take_bit_of_2 [simp]: \take_bit n 2 = of_bool (2 \ n) * 2\ using take_bit_of_exp [of n 1] by simp lemma take_bit_of_mask: \take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\ by (simp add: take_bit_eq_mod mask_mod_exp) lemma push_bit_eq_0_iff [simp]: "push_bit n a = 0 \ a = 0" by (simp add: push_bit_eq_mult) lemma take_bit_add: "take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)" by (simp add: take_bit_eq_mod mod_simps) lemma take_bit_of_1_eq_0_iff [simp]: "take_bit n 1 = 0 \ n = 0" by (simp add: take_bit_eq_mod) lemma take_bit_Suc_1 [simp]: \take_bit (Suc n) 1 = 1\ by (simp add: take_bit_Suc) lemma take_bit_Suc_bit0 [simp]: \take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\ by (simp add: take_bit_Suc numeral_Bit0_div_2) lemma take_bit_Suc_bit1 [simp]: \take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\ by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd) lemma take_bit_numeral_1 [simp]: \take_bit (numeral l) 1 = 1\ by (simp add: take_bit_rec [of \numeral l\ 1]) lemma take_bit_numeral_bit0 [simp]: \take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\ by (simp add: take_bit_rec numeral_Bit0_div_2) lemma take_bit_numeral_bit1 [simp]: \take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\ by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd) lemma drop_bit_Suc_bit0 [simp]: \drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\ by (simp add: drop_bit_Suc numeral_Bit0_div_2) lemma drop_bit_Suc_bit1 [simp]: \drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\ by (simp add: drop_bit_Suc numeral_Bit1_div_2) lemma drop_bit_numeral_bit0 [simp]: \drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\ by (simp add: drop_bit_rec numeral_Bit0_div_2) lemma drop_bit_numeral_bit1 [simp]: \drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\ by (simp add: drop_bit_rec numeral_Bit1_div_2) lemma drop_bit_of_nat: "drop_bit n (of_nat m) = of_nat (drop_bit n m)" by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) lemma bit_of_nat_iff_bit [simp]: \bit (of_nat m) n \ bit m n\ proof - have \even (m div 2 ^ n) \ even (of_nat (m div 2 ^ n))\ by simp also have \of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\ by (simp add: of_nat_div) finally show ?thesis by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd) qed lemma of_nat_drop_bit: \of_nat (drop_bit m n) = drop_bit m (of_nat n)\ by (simp add: drop_bit_eq_div semiring_bit_shifts_class.drop_bit_eq_div of_nat_div) lemma bit_push_bit_iff_of_nat_iff: \bit (push_bit m (of_nat r)) n \ m \ n \ bit (of_nat r) (n - m)\ by (auto simp add: bit_push_bit_iff) end instance nat :: unique_euclidean_semiring_with_bit_shifts .. instance int :: unique_euclidean_semiring_with_bit_shifts .. lemma bit_nat_iff: \bit (nat k) n \ k \ 0 \ bit k n\ proof (cases \k \ 0\) case True moreover define m where \m = nat k\ ultimately have \k = int m\ by simp then show ?thesis by (auto intro: ccontr) next case False then show ?thesis by simp qed lemma push_bit_nat_eq: \push_bit n (nat k) = nat (push_bit n k)\ by (cases \k \ 0\) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2) lemma drop_bit_nat_eq: \drop_bit n (nat k) = nat (drop_bit n k)\ apply (cases \k \ 0\) apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le) apply (simp add: divide_int_def) done lemma take_bit_nat_eq: \take_bit n (nat k) = nat (take_bit n k)\ if \k \ 0\ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma nat_take_bit_eq: \nat (take_bit n k) = take_bit n (nat k)\ if \k \ 0\ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma not_exp_less_eq_0_int [simp]: \\ 2 ^ n \ (0::int)\ by (simp add: power_le_zero_eq) lemma half_nonnegative_int_iff [simp]: \k div 2 \ 0 \ k \ 0\ for k :: int proof (cases \k \ 0\) case True then show ?thesis by (auto simp add: divide_int_def sgn_1_pos) next case False then show ?thesis apply (auto simp add: divide_int_def not_le elim!: evenE) apply (simp only: minus_mult_right) apply (subst nat_mult_distrib) apply simp_all done qed lemma half_negative_int_iff [simp]: \k div 2 < 0 \ k < 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma push_bit_of_Suc_0 [simp]: "push_bit n (Suc 0) = 2 ^ n" using push_bit_of_1 [where ?'a = nat] by simp lemma take_bit_of_Suc_0 [simp]: "take_bit n (Suc 0) = of_bool (0 < n)" using take_bit_of_1 [where ?'a = nat] by simp lemma drop_bit_of_Suc_0 [simp]: "drop_bit n (Suc 0) = of_bool (n = 0)" using drop_bit_of_1 [where ?'a = nat] by simp lemma push_bit_minus_one: "push_bit n (- 1 :: int) = - (2 ^ n)" by (simp add: push_bit_eq_mult) lemma minus_1_div_exp_eq_int: \- 1 div (2 :: int) ^ n = - 1\ by (induction n) (use div_exp_eq [symmetric, of \- 1 :: int\ 1] in \simp_all add: ac_simps\) lemma drop_bit_minus_one [simp]: \drop_bit n (- 1 :: int) = - 1\ by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int) lemma take_bit_Suc_from_most: \take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\ for k :: int by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq) lemma take_bit_minus: \take_bit n (- take_bit n k) = take_bit n (- k)\ for k :: int by (simp add: take_bit_eq_mod mod_minus_eq) lemma take_bit_diff: \take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\ for k l :: int by (simp add: take_bit_eq_mod mod_diff_eq) lemma bit_imp_take_bit_positive: \0 < take_bit m k\ if \n < m\ and \bit k n\ for k :: int proof (rule ccontr) assume \\ 0 < take_bit m k\ then have \take_bit m k = 0\ by (auto simp add: not_less intro: order_antisym) then have \bit (take_bit m k) n = bit 0 n\ by simp with that show False by (simp add: bit_take_bit_iff) qed lemma take_bit_mult: \take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\ for k l :: int by (simp add: take_bit_eq_mod mod_mult_eq) lemma (in ring_1) of_nat_nat_take_bit_eq [simp]: \of_nat (nat (take_bit n k)) = of_int (take_bit n k)\ by simp lemma take_bit_minus_small_eq: \take_bit n (- k) = 2 ^ n - k\ if \0 < k\ \k \ 2 ^ n\ for k :: int proof - define m where \m = nat k\ with that have \k = int m\ and \0 < m\ and \m \ 2 ^ n\ by simp_all have \(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\ using \0 < m\ by simp then have \int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\ by simp then have \(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\ using \m \ 2 ^ n\ by (simp only: of_nat_mod of_nat_diff) simp with \k = int m\ have \(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\ by simp then show ?thesis by (simp add: take_bit_eq_mod) qed lemma drop_bit_push_bit_int: \drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\ for k :: int by (cases \m \ n\) (auto simp add: mult.left_commute [of _ \2 ^ n\] mult.commute [of _ \2 ^ n\] mult.assoc mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add) lemma push_bit_nonnegative_int_iff [simp]: \push_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: push_bit_eq_mult zero_le_mult_iff) lemma push_bit_negative_int_iff [simp]: \push_bit n k < 0 \ k < 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma drop_bit_nonnegative_int_iff [simp]: \drop_bit n k \ 0 \ k \ 0\ for k :: int by (induction n) (simp_all add: drop_bit_Suc drop_bit_half) lemma drop_bit_negative_int_iff [simp]: \drop_bit n k < 0 \ k < 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) code_identifier code_module Parity \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/ROOT b/src/HOL/ROOT --- a/src/HOL/ROOT +++ b/src/HOL/ROOT @@ -1,1171 +1,1170 @@ chapter HOL session HOL (main) = Pure + description " Classical Higher-order Logic. " options [strict_facts] sessions Tools theories Main (global) Complex_Main (global) document_files "root.bib" "root.tex" session "HOL-Examples" in Examples = HOL + description " Notable Examples in Isabelle/HOL. " sessions "HOL-Library" theories Adhoc_Overloading_Examples Ackermann Cantor Coherent Commands Drinker Groebner_Examples Iff_Oracle Induction_Schema Knaster_Tarski "ML" Peirce Records Seq document_files "root.bib" "root.tex" session "HOL-Proofs" (timing) in Proofs = Pure + description " HOL-Main with explicit proof terms. " options [quick_and_dirty = false, record_proofs = 2, parallel_limit = 500] sessions "HOL-Library" theories "HOL-Library.Realizers" session "HOL-Library" (main timing) in Library = HOL + description " Classical Higher-order Logic -- batteries included. " theories Library (*conflicting type class instantiations and dependent applications*) Finite_Lattice List_Lexorder List_Lenlexorder Prefix_Order Product_Lexorder Product_Order Subseq_Order (*conflicting syntax*) Datatype_Records (*data refinements and dependent applications*) AList_Mapping Code_Binary_Nat Code_Prolog Code_Real_Approx_By_Float Code_Target_Numeral DAList DAList_Multiset RBT_Mapping RBT_Set (*printing modifications*) OptionalSugar (*prototypic tools*) Predicate_Compile_Quickcheck (*legacy tools*) Old_Datatype Old_Recdef Realizers Refute document_files "root.bib" "root.tex" session "HOL-Analysis" (main timing) in Analysis = HOL + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Library" "HOL-Computational_Algebra" theories Analysis document_files "root.tex" "root.bib" session "HOL-Complex_Analysis" (main timing) in Complex_Analysis = "HOL-Analysis" + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] theories Complex_Analysis document_files "root.tex" "root.bib" session "HOL-Analysis-ex" in "Analysis/ex" = "HOL-Analysis" + theories Approximations Metric_Arith_Examples session "HOL-Homology" (timing) in Homology = "HOL-Analysis" + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Algebra" theories Homology document_files "root.tex" session "HOL-Computational_Algebra" (main timing) in "Computational_Algebra" = "HOL-Library" + theories Computational_Algebra (*conflicting type class instantiations and dependent applications*) Field_as_Ring session "HOL-Real_Asymp" in Real_Asymp = HOL + sessions "HOL-Decision_Procs" theories Real_Asymp Real_Asymp_Approx Real_Asymp_Examples session "HOL-Real_Asymp-Manual" in "Real_Asymp/Manual" = "HOL-Real_Asymp" + theories Real_Asymp_Doc document_files (in "~~/src/Doc") "iman.sty" "extra.sty" "isar.sty" document_files "root.tex" "style.sty" session "HOL-Hahn_Banach" in Hahn_Banach = HOL + description " Author: Gertrud Bauer, TU Munich The Hahn-Banach theorem for real vector spaces. This is the proof of the Hahn-Banach theorem for real vectorspaces, following H. Heuser, Funktionalanalysis, p. 228 -232. The Hahn-Banach theorem is one of the fundamental theorems of functional analysis. It is a conclusion of Zorn's lemma. Two different formaulations of the theorem are presented, one for general real vectorspaces and its application to normed vectorspaces. The theorem says, that every continous linearform, defined on arbitrary subspaces (not only one-dimensional subspaces), can be extended to a continous linearform on the whole vectorspace. " sessions "HOL-Analysis" theories Hahn_Banach document_files "root.bib" "root.tex" session "HOL-Induct" in Induct = HOL + description " Examples of (Co)Inductive Definitions. Comb proves the Church-Rosser theorem for combinators (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR396-lcp-generic-automatic-proof-tools.ps.gz). Mutil is the famous Mutilated Chess Board problem (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR394-lcp-mutilated-chess-board.dvi.gz). PropLog proves the completeness of a formalization of propositional logic (see http://www.cl.cam.ac.uk/Research/Reports/TR312-lcp-set-II.ps.gz). Exp demonstrates the use of iterated inductive definitions to reason about mutually recursive relations. " sessions "HOL-Library" theories [quick_and_dirty] Common_Patterns theories Nested_Datatype QuoDataType QuoNestedDataType Term SList ABexp Infinitely_Branching_Tree Ordinals Sigma_Algebra Comb PropLog Com document_files "root.tex" session "HOL-IMP" (timing) in IMP = "HOL-Library" + options [document_variants = document] theories BExp ASM Finite_Reachable Denotational Compiler2 Poly_Types Sec_Typing Sec_TypingT Def_Init_Big Def_Init_Small Fold Live Live_True Hoare_Examples Hoare_Sound_Complete VCG Hoare_Total VCG_Total_EX VCG_Total_EX2 Collecting1 Collecting_Examples Abs_Int_Tests Abs_Int1_parity Abs_Int1_const Abs_Int3 Procs_Dyn_Vars_Dyn Procs_Stat_Vars_Dyn Procs_Stat_Vars_Stat C_like OO document_files "root.bib" "root.tex" session "HOL-IMPP" in IMPP = HOL + description \ Author: David von Oheimb Copyright 1999 TUM IMPP -- An imperative language with procedures. This is an extension of IMP with local variables and mutually recursive procedures. For documentation see "Hoare Logic for Mutual Recursion and Local Variables" (https://isabelle.in.tum.de/Bali/papers/FSTTCS99.html). \ theories EvenOdd session "HOL-Data_Structures" (timing) in Data_Structures = HOL + options [document_variants = document] sessions "HOL-Number_Theory" theories [document = false] Less_False theories Sorting Balance Tree_Map Interval_Tree AVL_Map AVL_Bal_Set AVL_Bal2_Set Height_Balanced_Tree RBT_Set2 RBT_Map Tree23_Map Tree23_of_List Tree234_Map Brother12_Map AA_Map Set2_Join_RBT Array_Braun Trie_Fun Trie_Map Tries_Binary Leftist_Heap Binomial_Heap document_files "root.tex" "root.bib" session "HOL-Import" in Import = HOL + theories HOL_Light_Maps theories [condition = HOL_LIGHT_BUNDLE] HOL_Light_Import session "HOL-Number_Theory" (main timing) in Number_Theory = "HOL-Computational_Algebra" + description " Fundamental Theorem of Arithmetic, Chinese Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, some lemmas for Quadratic Reciprocity. " sessions "HOL-Algebra" theories Number_Theory document_files "root.tex" session "HOL-Hoare" in Hoare = HOL + description " Verification of imperative programs (verification conditions are generated automatically from pre/post conditions and loop invariants). " theories Hoare document_files "root.bib" "root.tex" session "HOL-Hoare_Parallel" (timing) in Hoare_Parallel = HOL + description " Verification of shared-variable imperative programs a la Owicki-Gries. (verification conditions are generated automatically). " theories Hoare_Parallel document_files "root.bib" "root.tex" session "HOL-Codegenerator_Test" in Codegenerator_Test = "HOL-Number_Theory" + sessions "HOL-Data_Structures" "HOL-Examples" "HOL-Word" theories Generate Generate_Binary_Nat Generate_Target_Nat Generate_Efficient_Datastructures Code_Lazy_Test Code_Test_PolyML Code_Test_Scala theories [condition = ISABELLE_GHC] Code_Test_GHC theories [condition = ISABELLE_MLTON] Code_Test_MLton theories [condition = ISABELLE_OCAMLFIND] Code_Test_OCaml theories [condition = ISABELLE_SMLNJ] Code_Test_SMLNJ session "HOL-Metis_Examples" (timing) in Metis_Examples = "HOL-Library" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Testing Metis and Sledgehammer. " sessions "HOL-Decision_Procs" theories Abstraction Big_O Binary_Tree Clausification Message Proxies Tarski Trans_Closure Sets session "HOL-Nitpick_Examples" in Nitpick_Examples = HOL + description " Author: Jasmin Blanchette, TU Muenchen Copyright 2009 " options [kodkod_scala] sessions "HOL-Library" theories [quick_and_dirty] Nitpick_Examples session "HOL-Algebra" (main timing) in Algebra = "HOL-Computational_Algebra" + description " Author: Clemens Ballarin, started 24 September 1999, and many others The Isabelle Algebraic Library. " sessions "HOL-Cardinals" theories (* Orders and Lattices *) Galois_Connection (* Knaster-Tarski theorem and Galois connections *) (* Groups *) FiniteProduct (* Product operator for commutative groups *) Sylow (* Sylow's theorem *) Bij (* Automorphism Groups *) Multiplicative_Group Zassenhaus (* The Zassenhaus lemma *) (* Rings *) Divisibility (* Rings *) IntRing (* Ideals and residue classes *) UnivPoly (* Polynomials *) (* Main theory *) Algebra document_files "root.bib" "root.tex" session "HOL-Auth" (timing) in Auth = HOL + description " A new approach to verifying authentication protocols. " sessions "HOL-Library" directories "Smartcard" "Guard" theories Auth_Shared Auth_Public "Smartcard/Auth_Smartcard" "Guard/Auth_Guard_Shared" "Guard/Auth_Guard_Public" document_files "root.tex" session "HOL-UNITY" (timing) in UNITY = "HOL-Auth" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge Verifying security protocols using Chandy and Misra's UNITY formalism. " directories "Simple" "Comp" theories (*Basic meta-theory*) UNITY_Main (*Simple examples: no composition*) "Simple/Deadlock" "Simple/Common" "Simple/Network" "Simple/Token" "Simple/Channel" "Simple/Lift" "Simple/Mutex" "Simple/Reach" "Simple/Reachability" (*Verifying security protocols using UNITY*) "Simple/NSP_Bad" (*Example of composition*) "Comp/Handshake" (*Universal properties examples*) "Comp/Counter" "Comp/Counterc" "Comp/Priority" "Comp/TimerArray" "Comp/Progress" "Comp/Alloc" "Comp/AllocImpl" "Comp/Client" (*obsolete*) ELT document_files "root.tex" session "HOL-Unix" in Unix = HOL + options [print_mode = "no_brackets,no_type_brackets"] sessions "HOL-Library" theories Unix document_files "root.bib" "root.tex" session "HOL-ZF" in ZF = HOL + sessions "HOL-Library" theories MainZF Games document_files "root.tex" session "HOL-Imperative_HOL" (timing) in Imperative_HOL = HOL + options [print_mode = "iff,no_brackets"] sessions "HOL-Library" directories "ex" theories Imperative_HOL_ex document_files "root.bib" "root.tex" session "HOL-Decision_Procs" (timing) in Decision_Procs = "HOL-Algebra" + description " Various decision procedures, typically involving reflection. " directories "ex" theories Decision_Procs session "HOL-Proofs-ex" in "Proofs/ex" = "HOL-Proofs" + sessions "HOL-Examples" theories Hilbert_Classical Proof_Terms XML_Data session "HOL-Proofs-Extraction" (timing) in "Proofs/Extraction" = "HOL-Proofs" + description " Examples for program extraction in Higher-Order Logic. " options [quick_and_dirty = false] sessions "HOL-Computational_Algebra" theories Greatest_Common_Divisor Warshall Higman_Extraction Pigeonhole Euclid document_files "root.bib" "root.tex" session "HOL-Proofs-Lambda" (timing) in "Proofs/Lambda" = "HOL-Proofs" + description \ Lambda Calculus in de Bruijn's Notation. This session defines lambda-calculus terms with de Bruijn indixes and proves confluence of beta, eta and beta+eta. The paper "More Church-Rosser Proofs (in Isabelle/HOL)" describes the whole theory (see http://www.in.tum.de/~nipkow/pubs/jar2001.html). \ options [print_mode = "no_brackets", quick_and_dirty = false] sessions "HOL-Library" theories Eta StrongNorm Standardization WeakNorm document_files "root.bib" "root.tex" session "HOL-Prolog" in Prolog = HOL + description " Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) A bare-bones implementation of Lambda-Prolog. This is a simple exploratory implementation of Lambda-Prolog in HOL, including some minimal examples (in Test.thy) and a more typical example of a little functional language and its type system. " theories Test Type session "HOL-MicroJava" (timing) in MicroJava = HOL + description " Formalization of a fragment of Java, together with a corresponding virtual machine and a specification of its bytecode verifier and a lightweight bytecode verifier, including proofs of type-safety. " sessions "HOL-Library" "HOL-Eisbach" directories "BV" "Comp" "DFA" "J" "JVM" theories MicroJava document_files "introduction.tex" "root.bib" "root.tex" session "HOL-NanoJava" in NanoJava = HOL + description " Hoare Logic for a tiny fragment of Java. " theories Example document_files "root.bib" "root.tex" session "HOL-Bali" (timing) in Bali = HOL + sessions "HOL-Library" theories AxExample AxSound AxCompl Trans TypeSafe document_files "root.tex" session "HOL-IOA" in IOA = HOL + description \ Author: Tobias Nipkow and Konrad Slind and Olaf Müller Copyright 1994--1996 TU Muenchen The meta-theory of I/O-Automata in HOL. This formalization has been significantly changed and extended, see HOLCF/IOA. There are also the proofs of two communication protocols which formerly have been here. @inproceedings{Nipkow-Slind-IOA, author={Tobias Nipkow and Konrad Slind}, title={{I/O} Automata in {Isabelle/HOL}}, booktitle={Proc.\ TYPES Workshop 1994}, publisher=Springer, series=LNCS, note={To appear}} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/ioa.ps.gz and @inproceedings{Mueller-Nipkow, author={Olaf M\"uller and Tobias Nipkow}, title={Combining Model Checking and Deduction for {I/O}-Automata}, booktitle={Proc.\ TACAS Workshop}, organization={Aarhus University, BRICS report}, year=1995} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/tacas.dvi.gz \ theories Solve session "HOL-Lattice" in Lattice = HOL + description " Author: Markus Wenzel, TU Muenchen Basic theory of lattices and orders. " theories CompleteLattice document_files "root.tex" session "HOL-ex" (timing) in ex = "HOL-Number_Theory" + description " Miscellaneous examples for Higher-Order Logic. " sessions "HOL-Word" theories Antiquote Argo_Examples Arith_Examples Ballot BinEx Birthday_Paradox Bit_Lists Bubblesort CTL Cartouche_Examples Case_Product Chinese Classical Code_Binary_Nat_examples Code_Lazy_Demo Code_Timing Coercion_Examples Computations Conditional_Parametricity_Examples Cubic_Quartic Datatype_Record_Examples Dedekind_Real Erdoes_Szekeres Eval_Examples Executable_Relation Execute_Choice Functions Function_Growth Gauge_Integration Guess HarmonicSeries Hebrew Hex_Bin_Examples IArray_Examples Intuitionistic Join_Theory Lagrange List_to_Set_Comprehension_Examples LocaleTest2 MergeSort MonoidGroup Multiquote NatSum Normalization_by_Evaluation PER Parallel_Example Peano_Axioms Perm_Fragments PresburgerEx Primrec Pythagoras Quicksort Radix_Sort Reflection_Examples Refute_Examples Residue_Ring Rewrite_Examples SOS SOS_Cert Serbian Set_Comprehension_Pointfree_Examples Set_Theory Simproc_Tests Simps_Case_Conv_Examples Sketch_and_Explore Sorting_Algorithms_Examples Sqrt Sqrt_Script Sudoku Sum_of_Powers Tarski Termination ThreeDivides Transfer_Debug Transfer_Int_Nat Transitive_Closure_Table_Ex Tree23 Triangular_Numbers Unification While_Combinator_Example veriT_Preprocessing theories [skip_proofs = false] SAT_Examples Meson_Test session "HOL-Isar_Examples" in Isar_Examples = "HOL-Computational_Algebra" + description " Miscellaneous Isabelle/Isar examples. " options [quick_and_dirty] theories Structured_Statements Basic_Logic Expr_Compiler Fibonacci Group Group_Context Group_Notepad Hoare_Ex Mutilated_Checkerboard Puzzle Summation document_files "root.bib" "root.tex" session "HOL-Eisbach" in Eisbach = HOL + description \ The Eisbach proof method language and "match" method. \ sessions FOL "HOL-Analysis" theories Eisbach Tests Examples Examples_FOL Example_Metric session "HOL-SET_Protocol" (timing) in SET_Protocol = HOL + description " Verification of the SET Protocol. " sessions "HOL-Library" theories SET_Protocol document_files "root.tex" session "HOL-Matrix_LP" in Matrix_LP = HOL + description " Two-dimensional matrices and linear programming. " sessions "HOL-Library" directories "Compute_Oracle" theories Cplex document_files "root.tex" session "HOL-TLA" in TLA = HOL + description " Lamport's Temporal Logic of Actions. " theories TLA session "HOL-TLA-Inc" in "TLA/Inc" = "HOL-TLA" + theories Inc session "HOL-TLA-Buffer" in "TLA/Buffer" = "HOL-TLA" + theories DBuffer session "HOL-TLA-Memory" in "TLA/Memory" = "HOL-TLA" + theories MemoryImplementation session "HOL-TPTP" in TPTP = HOL + description " Author: Jasmin Blanchette, TU Muenchen Author: Nik Sultana, University of Cambridge Copyright 2011 TPTP-related extensions. " sessions "HOL-Library" theories ATP_Theory_Export MaSh_Eval TPTP_Interpret THF_Arith TPTP_Proof_Reconstruction theories ATP_Problem_Import session "HOL-Probability" (main timing) in "Probability" = "HOL-Analysis" + theories Probability document_files "root.tex" session "HOL-Probability-ex" (timing) in "Probability/ex" = "HOL-Probability" + theories Dining_Cryptographers Koepf_Duermuth_Countermeasure Measure_Not_CCC session "HOL-Nominal" in Nominal = HOL + sessions "HOL-Library" theories Nominal session "HOL-Nominal-Examples" (timing) in "Nominal/Examples" = "HOL-Nominal" + theories Class3 CK_Machine Compile Contexts Crary CR_Takahashi CR Fsub Height Lambda_mu Lam_Funs LocalWeakening Pattern SN SOS Standardization Support Type_Preservation Weakening W theories [quick_and_dirty] VC_Condition session "HOL-Cardinals" (timing) in Cardinals = HOL + description " Ordinals and Cardinals, Full Theories. " theories Cardinals Bounded_Set document_files "intro.tex" "root.tex" "root.bib" session "HOL-Datatype_Examples" (timing) in Datatype_Examples = "HOL-Library" + description " (Co)datatype Examples. " directories "Derivation_Trees" theories Compat Lambda_Term Process TreeFsetI "Derivation_Trees/Gram_Lang" "Derivation_Trees/Parallel_Composition" Koenig Lift_BNF Milner_Tofte Stream_Processor Cyclic_List Free_Idempotent_Monoid LDL TLList Misc_Codatatype Misc_Datatype Misc_Primcorec Misc_Primrec Datatype_Simproc_Tests session "HOL-Corec_Examples" (timing) in Corec_Examples = "HOL-Library" + description " Corecursion Examples. " directories "Tests" theories LFilter Paper_Examples Stream_Processor "Tests/Simple_Nesting" "Tests/Iterate_GPV" theories [quick_and_dirty] "Tests/GPV_Bare_Bones" "Tests/Merge_D" "Tests/Merge_Poly" "Tests/Misc_Mono" "Tests/Misc_Poly" "Tests/Small_Concrete" "Tests/Stream_Friends" "Tests/TLList_Friends" "Tests/Type_Class" session "HOL-Word" (main timing) in Word = HOL + sessions "HOL-Library" theories Word More_Word Word_Examples - Conversions document_files "root.bib" "root.tex" session "HOL-Statespace" in Statespace = HOL + theories [skip_proofs = false] StateSpaceEx document_files "root.tex" session "HOL-Nonstandard_Analysis" (timing) in Nonstandard_Analysis = "HOL-Computational_Algebra" + description " Nonstandard analysis. " theories Nonstandard_Analysis document_files "root.tex" session "HOL-Nonstandard_Analysis-Examples" (timing) in "Nonstandard_Analysis/Examples" = "HOL-Nonstandard_Analysis" + theories NSPrimes session "HOL-Mirabelle" in Mirabelle = HOL + theories Mirabelle_Test session "HOL-Mirabelle-ex" in "Mirabelle/ex" = "HOL-Mirabelle" + options [timeout = 60] theories Ex session "HOL-Word-SMT_Examples" (timing) in SMT_Examples = "HOL-Word" + options [quick_and_dirty] theories Boogie SMT_Examples SMT_Word_Examples SMT_Tests session "HOL-SPARK" in "SPARK" = "HOL-Word" + theories SPARK session "HOL-SPARK-Examples" in "SPARK/Examples" = "HOL-SPARK" + directories "Gcd" "Liseq" "RIPEMD-160" "Sqrt" theories "Gcd/Greatest_Common_Divisor" "Liseq/Longest_Increasing_Subsequence" "RIPEMD-160/F" "RIPEMD-160/Hash" "RIPEMD-160/K_L" "RIPEMD-160/K_R" "RIPEMD-160/R_L" "RIPEMD-160/Round" "RIPEMD-160/R_R" "RIPEMD-160/S_L" "RIPEMD-160/S_R" "Sqrt/Sqrt" export_files (in ".") "*:**.prv" session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" + options [show_question_marks = false] sessions "HOL-SPARK-Examples" theories Example_Verification VC_Principles Reference Complex_Types document_files "complex_types.ads" "complex_types_app.adb" "complex_types_app.ads" "Gcd.adb" "Gcd.ads" "intro.tex" "loop_invariant.adb" "loop_invariant.ads" "root.bib" "root.tex" "Simple_Gcd.adb" "Simple_Gcd.ads" session "HOL-Mutabelle" in Mutabelle = HOL + sessions "HOL-Library" theories MutabelleExtra session "HOL-Quickcheck_Examples" (timing) in Quickcheck_Examples = HOL + sessions "HOL-Library" theories Quickcheck_Examples Quickcheck_Lattice_Examples Completeness Quickcheck_Interfaces Quickcheck_Nesting_Example theories [condition = ISABELLE_GHC] Hotel_Example Quickcheck_Narrowing_Examples session "HOL-Quotient_Examples" (timing) in Quotient_Examples = "HOL-Algebra" + description " Author: Cezary Kaliszyk and Christian Urban " theories DList Quotient_FSet Quotient_Int Quotient_Message Lift_FSet Lift_Set Lift_Fun Quotient_Rat Lift_DList Int_Pow Lifting_Code_Dt_Test session "HOL-Predicate_Compile_Examples" (timing) in Predicate_Compile_Examples = HOL + sessions "HOL-Library" theories Examples Predicate_Compile_Tests Predicate_Compile_Quickcheck_Examples Specialisation_Examples IMP_1 IMP_2 (* FIXME since 21-Jul-2011 Hotel_Example_Small_Generator *) IMP_3 IMP_4 theories [condition = ISABELLE_SWIPL] Code_Prolog_Examples Context_Free_Grammar_Example Hotel_Example_Prolog Lambda_Example List_Examples theories [condition = ISABELLE_SWIPL, quick_and_dirty] Reg_Exp_Example session "HOL-Types_To_Sets" in Types_To_Sets = HOL + description " Experimental extension of Higher-Order Logic to allow translation of types to sets. " directories "Examples" theories Types_To_Sets "Examples/Prerequisites" "Examples/Finite" "Examples/T2_Spaces" "Examples/Unoverload_Def" "Examples/Linear_Algebra_On" session HOLCF (main timing) in HOLCF = HOL + description " Author: Franz Regensburger Author: Brian Huffman HOLCF -- a semantic extension of HOL by the LCF logic. " sessions "HOL-Library" theories HOLCF (global) document_files "root.tex" session "HOLCF-Tutorial" in "HOLCF/Tutorial" = HOLCF + theories Domain_ex Fixrec_ex New_Domain document_files "root.tex" session "HOLCF-Library" in "HOLCF/Library" = HOLCF + theories HOLCF_Library HOL_Cpo session "HOLCF-IMP" in "HOLCF/IMP" = HOLCF + description " IMP -- A WHILE-language and its Semantics. This is the HOLCF-based denotational semantics of a simple WHILE-language. " sessions "HOL-IMP" theories HoareEx document_files "isaverbatimwrite.sty" "root.tex" "root.bib" session "HOLCF-ex" in "HOLCF/ex" = "HOLCF-Library" + description " Miscellaneous examples for HOLCF. " theories Dnat Dagstuhl Focus_ex Fix2 Hoare Concurrency_Monad Loop Powerdomain_ex Domain_Proofs Letrec Pattern_Match session "HOLCF-FOCUS" in "HOLCF/FOCUS" = "HOLCF-Library" + description \ FOCUS: a theory of stream-processing functions Isabelle/HOLCF. For introductions to FOCUS, see "The Design of Distributed Systems - An Introduction to FOCUS" http://www4.in.tum.de/publ/html.php?e=2 "Specification and Refinement of a Buffer of Length One" http://www4.in.tum.de/publ/html.php?e=15 "Specification and Development of Interactive Systems: Focus on Streams, Interfaces, and Refinement" http://www4.in.tum.de/publ/html.php?e=321 \ theories Fstreams FOCUS Buffer_adm session IOA (timing) in "HOLCF/IOA" = HOLCF + description " Author: Olaf Mueller Copyright 1997 TU München A formalization of I/O automata in HOLCF. The distribution contains simulation relations, temporal logic, and an abstraction theory. Everything is based upon a domain-theoretic model of finite and infinite sequences. " theories Abstraction session "IOA-ABP" in "HOLCF/IOA/ABP" = IOA + description " Author: Olaf Mueller The Alternating Bit Protocol performed in I/O-Automata. " theories Correctness Spec session "IOA-NTP" in "HOLCF/IOA/NTP" = IOA + description " Author: Tobias Nipkow & Konrad Slind A network transmission protocol, performed in the I/O automata formalization by Olaf Mueller. " theories Correctness session "IOA-Storage" in "HOLCF/IOA/Storage" = IOA + description " Author: Olaf Mueller Memory storage case study. " theories Correctness session "IOA-ex" in "HOLCF/IOA/ex" = IOA + description " Author: Olaf Mueller " theories TrivEx TrivEx2 diff --git a/src/HOL/SPARK/Examples/RIPEMD-160/F.thy b/src/HOL/SPARK/Examples/RIPEMD-160/F.thy --- a/src/HOL/SPARK/Examples/RIPEMD-160/F.thy +++ b/src/HOL/SPARK/Examples/RIPEMD-160/F.thy @@ -1,66 +1,67 @@ (* Title: HOL/SPARK/Examples/RIPEMD-160/F.thy Author: Fabian Immler, TU Muenchen Verification of the RIPEMD-160 hash function *) theory F imports RMD_Specification begin spark_open \rmd/f\ spark_vc function_f_2 using assms by simp_all spark_vc function_f_3 using assms by simp_all spark_vc function_f_4 using assms by simp_all spark_vc function_f_5 using assms by simp_all spark_vc function_f_6 proof - from H8 have "nat j <= 15" by simp with assms show ?thesis - by (simp add: f_def bwsimps int_word_uint) + by (simp add: f_def bwsimps int_word_uint take_bit_int_eq_self) qed spark_vc function_f_7 proof - from H7 have "16 <= nat j" by simp moreover from H8 have "nat j <= 31" by simp ultimately show ?thesis using assms - by (simp add: f_def bwsimps int_word_uint) + by (simp only: f_def bwsimps int_word_uint) + (simp add: take_bit_int_eq_self take_bit_not_eq_mask_diff mask_eq_exp_minus_1) qed spark_vc function_f_8 proof - from H7 have "32 <= nat j" by simp moreover from H8 have "nat j <= 47" by simp ultimately show ?thesis using assms - by (simp add: f_def bwsimps int_word_uint) + by (simp only: f_def bwsimps int_word_uint) (simp add: take_bit_int_eq_self take_bit_not_eq_mask_diff mask_eq_exp_minus_1) qed spark_vc function_f_9 proof - from H7 have "48 <= nat j" by simp moreover from H8 have "nat j <= 63" by simp ultimately show ?thesis using assms - by (simp add: f_def bwsimps int_word_uint) + by (simp only: f_def bwsimps int_word_uint) (simp add: take_bit_int_eq_self take_bit_not_eq_mask_diff mask_eq_exp_minus_1) qed spark_vc function_f_10 proof - from H2 have "nat j <= 79" by simp moreover from H12 have "64 <= nat j" by simp ultimately show ?thesis using assms - by (simp add: f_def bwsimps int_word_uint) + by (simp only: f_def bwsimps int_word_uint) (simp add: take_bit_int_eq_self take_bit_not_eq_mask_diff mask_eq_exp_minus_1) qed spark_end end diff --git a/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy b/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy --- a/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy +++ b/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy @@ -1,100 +1,100 @@ (* Title: HOL/SPARK/Examples/RIPEMD-160/Hash.thy Author: Fabian Immler, TU Muenchen Verification of the RIPEMD-160 hash function *) theory Hash imports RMD_Specification begin spark_open \rmd/hash\ abbreviation from_chain :: "chain \ RMD.chain" where "from_chain c \ ( word_of_int (h0 c), word_of_int (h1 c), word_of_int (h2 c), word_of_int (h3 c), word_of_int (h4 c))" abbreviation to_chain :: "RMD.chain \ chain" where "to_chain c \ (let (h0, h1, h2, h3, h4) = c in (|h0 = uint h0, h1 = uint h1, h2 = uint h2, h3 = uint h3, h4 = uint h4|))" abbreviation round' :: "chain \ block \ chain" where "round' c b == to_chain (round (\n. word_of_int (b (int n))) (from_chain c))" abbreviation rounds' :: "chain \ int \ message \ chain" where "rounds' h i X == to_chain (rounds (\n. \m. word_of_int (X (int n) (int m))) (from_chain h) (nat i))" abbreviation rmd_hash :: "message \ int \ chain" where "rmd_hash X i == to_chain (rmd (\n. \m. word_of_int (X (int n) (int m))) (nat i))" spark_proof_functions round_spec = round' rounds = rounds' rmd_hash = rmd_hash spark_vc function_hash_12 using H1 H6 by (simp add: rounds_def rmd_body_def round_def h_0_def h0_0_def h1_0_def h2_0_def h3_0_def h4_0_def) lemma rounds_step: assumes "0 <= i" shows "rounds X b (Suc i) = round (X i) (rounds X b i)" by (simp add: rounds_def rmd_body_def) lemma from_to_id: "from_chain (to_chain C) = C" proof (cases C) fix a b c d e f::word32 assume "C = (a, b, c, d, e)" - thus ?thesis by (cases a) simp + thus ?thesis by (cases a) (simp add: take_bit_nat_eq_self) qed lemma steps_to_steps': "round X (foldl a b c) = round X (from_chain (to_chain (foldl a b c)))" unfolding from_to_id .. lemma rounds'_step: assumes "0 <= i" shows "rounds' c (i + 1) x = round' (rounds' c i x) (x i)" proof - have makesuc: "nat (i + 1) = Suc (nat i)" using assms by simp show ?thesis using assms by (simp add: makesuc rounds_def rmd_body_def steps_to_steps') qed spark_vc function_hash_13 proof - have loop_suc: "loop__1__i + 2 = (loop__1__i + 1) + 1" by simp have "0 <= loop__1__i + 1" using \0 <= loop__1__i\ by simp show ?thesis unfolding loop_suc unfolding rounds'_step[OF \0 <= loop__1__i + 1\] unfolding H1[symmetric] unfolding H18 .. qed spark_vc function_hash_17 unfolding rmd_def H1 rounds_def .. spark_end end diff --git a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy --- a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy +++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy @@ -1,466 +1,468 @@ (* Title: HOL/SPARK/Examples/RIPEMD-160/Round.thy Author: Fabian Immler, TU Muenchen Verification of the RIPEMD-160 hash function *) theory Round imports RMD_Specification begin spark_open \rmd/round\ abbreviation from_chain :: "chain \ RMD.chain" where "from_chain c \ ( word_of_int (h0 c), word_of_int (h1 c), word_of_int (h2 c), word_of_int (h3 c), word_of_int (h4 c))" abbreviation from_chain_pair :: "chain_pair \ RMD.chain \ RMD.chain" where "from_chain_pair cc \ ( from_chain (left cc), from_chain (right cc))" abbreviation to_chain :: "RMD.chain \ chain" where "to_chain c \ (let (h0, h1, h2, h3, h4) = c in (|h0 = uint h0, h1 = uint h1, h2 = uint h2, h3 = uint h3, h4 = uint h4|))" abbreviation to_chain_pair :: "RMD.chain \ RMD.chain \ chain_pair" where "to_chain_pair c == (let (c1, c2) = c in (| left = to_chain c1, right = to_chain c2 |))" abbreviation steps' :: "chain_pair \ int \ block \ chain_pair" where "steps' cc i b == to_chain_pair (steps (\n. word_of_int (b (int n))) (from_chain_pair cc) (nat i))" abbreviation round_spec :: "chain \ block \ chain" where "round_spec c b == to_chain (round (\n. word_of_int (b (int n))) (from_chain c))" spark_proof_functions steps = steps' round_spec = round_spec lemma uint_word_of_int_id: assumes "0 <= (x::int)" assumes "x <= 4294967295" shows"uint(word_of_int x::word32) = x" unfolding int_word_uint using assms by simp lemma steps_step: "steps X cc (Suc i) = step_both X (steps X cc i) i" unfolding steps_def by (induct i) simp_all lemma from_to_id: "from_chain_pair (to_chain_pair CC) = CC" proof (cases CC) fix a::RMD.chain fix b c d e f::word32 assume "CC = (a, b, c, d, e, f)" thus ?thesis by (cases a) simp qed lemma steps_to_steps': "F A (steps X cc i) B = F A (from_chain_pair (to_chain_pair (steps X cc i))) B" unfolding from_to_id .. lemma steps'_step: assumes "0 <= i" shows "steps' cc (i + 1) X = to_chain_pair ( step_both (\n. word_of_int (X (int n))) (from_chain_pair (steps' cc i X)) (nat i))" proof - have "nat (i + 1) = Suc (nat i)" using assms by simp show ?thesis unfolding \nat (i + 1) = Suc (nat i)\ steps_step steps_to_steps' .. qed lemma step_from_hyp: assumes step_hyp: "\left = \h0 = a, h1 = b, h2 = c, h3 = d, h4 = e\, right = \h0 = a', h1 = b', h2 = c', h3 = d', h4 = e'\\ = steps' (\left = \h0 = a_0, h1 = b_0, h2 = c_0, h3 = d_0, h4 = e_0\, right = \h0 = a_0, h1 = b_0, h2 = c_0, h3 = d_0, h4 = e_0\\) j x" assumes "a <= 4294967295" (is "_ <= ?M") assumes "b <= ?M" and "c <= ?M" and "d <= ?M" and "e <= ?M" assumes "a' <= ?M" and "b' <= ?M" and "c' <= ?M" and "d' <= ?M" and "e' <= ?M" assumes "0 <= a " and "0 <= b " and "0 <= c " and "0 <= d " and "0 <= e " assumes "0 <= a'" and "0 <= b'" and "0 <= c'" and "0 <= d'" and "0 <= e'" assumes "0 <= x (r_l_spec j)" and "x (r_l_spec j) <= ?M" assumes "0 <= x (r_r_spec j)" and "x (r_r_spec j) <= ?M" assumes "0 <= j" and "j <= 79" shows "\left = \h0 = e, h1 = (rotate_left (s_l_spec j) ((((a + f_spec j b c d) mod 4294967296 + x (r_l_spec j)) mod 4294967296 + k_l_spec j) mod 4294967296) + e) mod 4294967296, h2 = b, h3 = rotate_left 10 c, h4 = d\, right = \h0 = e', h1 = (rotate_left (s_r_spec j) ((((a' + f_spec (79 - j) b' c' d') mod 4294967296 + x (r_r_spec j)) mod 4294967296 + k_r_spec j) mod 4294967296) + e') mod 4294967296, h2 = b', h3 = rotate_left 10 c', h4 = d'\\ = steps' (\left = \h0 = a_0, h1 = b_0, h2 = c_0, h3 = d_0, h4 = e_0\, right = \h0 = a_0, h1 = b_0, h2 = c_0, h3 = d_0, h4 = e_0\\) (j + 1) x" using step_hyp proof - let ?MM = 4294967296 have AL: "uint(word_of_int e::word32) = e" by (rule uint_word_of_int_id[OF \0 <= e\ \e <= ?M\]) have CL: "uint(word_of_int b::word32) = b" by (rule uint_word_of_int_id[OF \0 <= b\ \b <= ?M\]) have DL: "True" .. have EL: "uint(word_of_int d::word32) = d" by (rule uint_word_of_int_id[OF \0 <= d\ \d <= ?M\]) have AR: "uint(word_of_int e'::word32) = e'" by (rule uint_word_of_int_id[OF \0 <= e'\ \e' <= ?M\]) have CR: "uint(word_of_int b'::word32) = b'" by (rule uint_word_of_int_id[OF \0 <= b'\ \b' <= ?M\]) have DR: "True" .. have ER: "uint(word_of_int d'::word32) = d'" by (rule uint_word_of_int_id[OF \0 <= d'\ \d' <= ?M\]) have BL: "(uint (word_rotl (s (nat j)) ((word_of_int::int\word32) ((((a + f_spec j b c d) mod ?MM + x (r_l_spec j)) mod ?MM + k_l_spec j) mod ?MM))) + e) mod ?MM = uint (word_rotl (s (nat j)) (word_of_int a + f (nat j) (word_of_int b) (word_of_int c) (word_of_int d) + word_of_int (x (r_l_spec j)) + K (nat j)) + word_of_int e)" (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _") proof - have "a mod ?MM = a" using \0 <= a\ \a <= ?M\ by simp have "?X mod ?MM = ?X" using \0 <= ?X\ \?X <= ?M\ by simp have "e mod ?MM = e" using \0 <= e\ \e <= ?M\ by simp have "(?MM::int) = 2 ^ LENGTH(32)" by simp show ?thesis unfolding word_add_def uint_word_of_int_id[OF \0 <= a\ \a <= ?M\] uint_word_of_int_id[OF \0 <= ?X\ \?X <= ?M\] int_word_uint unfolding \?MM = 2 ^ LENGTH(32)\ unfolding word_uint.Abs_norm by (simp add: \a mod ?MM = a\ \e mod ?MM = e\ \?X mod ?MM = ?X\) qed have BR: "(uint (word_rotl (s' (nat j)) ((word_of_int::int\word32) ((((a' + f_spec (79 - j) b' c' d') mod ?MM + x (r_r_spec j)) mod ?MM + k_r_spec j) mod ?MM))) + e') mod ?MM = uint (word_rotl (s' (nat j)) (word_of_int a' + f (79 - nat j) (word_of_int b') (word_of_int c') (word_of_int d') + word_of_int (x (r_r_spec j)) + K' (nat j)) + word_of_int e')" (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _") proof - have "a' mod ?MM = a'" using \0 <= a'\ \a' <= ?M\ by simp have "?X mod ?MM = ?X" using \0 <= ?X\ \?X <= ?M\ by simp have "e' mod ?MM = e'" using \0 <= e'\ \e' <= ?M\ by simp have "(?MM::int) = 2 ^ LENGTH(32)" by simp have nat_transfer: "79 - nat j = nat (79 - j)" using nat_diff_distrib \0 <= j\ \j <= 79\ by simp show ?thesis unfolding word_add_def uint_word_of_int_id[OF \0 <= a'\ \a' <= ?M\] uint_word_of_int_id[OF \0 <= ?X\ \?X <= ?M\] int_word_uint nat_transfer unfolding \?MM = 2 ^ LENGTH(32)\ unfolding word_uint.Abs_norm by (simp add: \a' mod ?MM = a'\ \e' mod ?MM = e'\ \?X mod ?MM = ?X\) qed show ?thesis unfolding steps'_step[OF \0 <= j\] step_hyp[symmetric] step_both_def step_r_def step_l_def - by (simp add: AL BL CL DL EL AR BR CR DR ER) + using AL CL EL AR CR ER + by (simp add: BL DL BR DR take_bit_int_eq_self_iff take_bit_int_eq_self) qed spark_vc procedure_round_61 proof - let ?M = "4294967295::int" have step_hyp: "\left = \h0 = ca, h1 = cb, h2 = cc, h3 = cd, h4 = ce\, right = \h0 = ca, h1 = cb, h2 = cc, h3 = cd, h4 = ce\\ = steps' (\left = \h0 = ca, h1 = cb, h2 = cc, h3 = cd, h4 = ce\, right = \h0 = ca, h1 = cb, h2 = cc, h3 = cd, h4 = ce\\) 0 x" unfolding steps_def - by (simp add: + using uint_word_of_int_id[OF \0 <= ca\ \ca <= ?M\] uint_word_of_int_id[OF \0 <= cb\ \cb <= ?M\] uint_word_of_int_id[OF \0 <= cc\ \cc <= ?M\] uint_word_of_int_id[OF \0 <= cd\ \cd <= ?M\] - uint_word_of_int_id[OF \0 <= ce\ \ce <= ?M\]) + uint_word_of_int_id[OF \0 <= ce\ \ce <= ?M\] + by (simp add: take_bit_int_eq_self_iff take_bit_int_eq_self) let ?rotate_arg_l = "((((ca + f 0 cb cc cd) mod 4294967296 + x (r_l 0)) mod 4294967296 + k_l 0) mod 4294967296)" let ?rotate_arg_r = "((((ca + f 79 cb cc cd) mod 4294967296 + x (r_r 0)) mod 4294967296 + k_r 0) mod 4294967296)" note returns = \wordops__rotate (s_l 0) ?rotate_arg_l = rotate_left (s_l 0) ?rotate_arg_l\ \wordops__rotate (s_r 0) ?rotate_arg_r = rotate_left (s_r 0) ?rotate_arg_r\ \wordops__rotate 10 cc = rotate_left 10 cc\ \f 0 cb cc cd = f_spec 0 cb cc cd\ \f 79 cb cc cd = f_spec 79 cb cc cd\ \k_l 0 = k_l_spec 0\ \k_r 0 = k_r_spec 0\ \r_l 0 = r_l_spec 0\ \r_r 0 = r_r_spec 0\ \s_l 0 = s_l_spec 0\ \s_r 0 = s_r_spec 0\ note x_borders = \\i. 0 \ i \ i \ 15 \ 0 \ x i \ x i \ ?M\ from \0 <= r_l 0\ \r_l 0 <= 15\ x_borders have "0 \ x (r_l 0)" by blast hence x_lower: "0 <= x (r_l_spec 0)" unfolding returns . from \0 <= r_l 0\ \r_l 0 <= 15\ x_borders have "x (r_l 0) <= ?M" by blast hence x_upper: "x (r_l_spec 0) <= ?M" unfolding returns . from \0 <= r_r 0\ \r_r 0 <= 15\ x_borders have "0 \ x (r_r 0)" by blast hence x_lower': "0 <= x (r_r_spec 0)" unfolding returns . from \0 <= r_r 0\ \r_r 0 <= 15\ x_borders have "x (r_r 0) <= ?M" by blast hence x_upper': "x (r_r_spec 0) <= ?M" unfolding returns . have "0 <= (0::int)" by simp have "0 <= (79::int)" by simp note step_from_hyp [OF step_hyp H2 H4 H6 H8 H10 H2 H4 H6 H8 H10 \ \upper bounds\ H1 H3 H5 H7 H9 H1 H3 H5 H7 H9 \ \lower bounds\ ] from this[OF x_lower x_upper x_lower' x_upper' \0 <= 0\ \0 <= 79\] \0 \ ca\ \0 \ ce\ x_lower x_lower' show ?thesis unfolding returns(1) returns(2) unfolding returns by (simp del: mod_pos_pos_trivial mod_neg_neg_trivial) qed spark_vc procedure_round_62 proof - let ?M = "4294967295::int" let ?rotate_arg_l = "((((cla + f (loop__1__j + 1) clb clc cld) mod 4294967296 + x (r_l (loop__1__j + 1))) mod 4294967296 + k_l (loop__1__j + 1)) mod 4294967296)" let ?rotate_arg_r = "((((cra + f (79 - (loop__1__j + 1)) crb crc crd) mod 4294967296 + x (r_r (loop__1__j + 1))) mod 4294967296 + k_r (loop__1__j + 1)) mod 4294967296)" have s: "78 - loop__1__j = (79 - (loop__1__j + 1))" by simp note returns = \wordops__rotate (s_l (loop__1__j + 1)) ?rotate_arg_l = rotate_left (s_l (loop__1__j + 1)) ?rotate_arg_l\ \wordops__rotate (s_r (loop__1__j + 1)) ?rotate_arg_r = rotate_left (s_r (loop__1__j + 1)) ?rotate_arg_r\ \f (loop__1__j + 1) clb clc cld = f_spec (loop__1__j + 1) clb clc cld\ \f (78 - loop__1__j) crb crc crd = f_spec (78 - loop__1__j) crb crc crd\[simplified s] \wordops__rotate 10 clc = rotate_left 10 clc\ \wordops__rotate 10 crc = rotate_left 10 crc\ \k_l (loop__1__j + 1) = k_l_spec (loop__1__j + 1)\ \k_r (loop__1__j + 1) = k_r_spec (loop__1__j + 1)\ \r_l (loop__1__j + 1) = r_l_spec (loop__1__j + 1)\ \r_r (loop__1__j + 1) = r_r_spec (loop__1__j + 1)\ \s_l (loop__1__j + 1) = s_l_spec (loop__1__j + 1)\ \s_r (loop__1__j + 1) = s_r_spec (loop__1__j + 1)\ note x_borders = \\i. 0 \ i \ i \ 15 \ 0 \ x i \ x i \ ?M\ from \0 <= r_l (loop__1__j + 1)\ \r_l (loop__1__j + 1) <= 15\ x_borders have "0 \ x (r_l (loop__1__j + 1))" by blast hence x_lower: "0 <= x (r_l_spec (loop__1__j + 1))" unfolding returns . from \0 <= r_l (loop__1__j + 1)\ \r_l (loop__1__j + 1) <= 15\ x_borders have "x (r_l (loop__1__j + 1)) <= ?M" by blast hence x_upper: "x (r_l_spec (loop__1__j + 1)) <= ?M" unfolding returns . from \0 <= r_r (loop__1__j + 1)\ \r_r (loop__1__j + 1) <= 15\ x_borders have "0 \ x (r_r (loop__1__j + 1))" by blast hence x_lower': "0 <= x (r_r_spec (loop__1__j + 1))" unfolding returns . from \0 <= r_r (loop__1__j + 1)\ \r_r (loop__1__j + 1) <= 15\ x_borders have "x (r_r (loop__1__j + 1)) <= ?M" by blast hence x_upper': "x (r_r_spec (loop__1__j + 1)) <= ?M" unfolding returns . from \0 <= loop__1__j\ have "0 <= loop__1__j + 1" by simp from \loop__1__j <= 78\ have "loop__1__j + 1 <= 79" by simp have "loop__1__j + 1 + 1 = loop__1__j + 2" by simp note step_from_hyp[OF H1 \cla <= ?M\ \clb <= ?M\ \clc <= ?M\ \cld <= ?M\ \cle <= ?M\ \cra <= ?M\ \crb <= ?M\ \crc <= ?M\ \crd <= ?M\ \cre <= ?M\ \0 <= cla\ \0 <= clb\ \0 <= clc\ \0 <= cld\ \0 <= cle\ \0 <= cra\ \0 <= crb\ \0 <= crc\ \0 <= crd\ \0 <= cre\] from this[OF x_lower x_upper x_lower' x_upper' \0 <= loop__1__j + 1\ \loop__1__j + 1 <= 79\] \0 \ cla\ \0 \ cle\ \0 \ cra\ \0 \ cre\ x_lower x_lower' show ?thesis unfolding \loop__1__j + 1 + 1 = loop__1__j + 2\ unfolding returns(1) returns(2) unfolding returns by (simp del: mod_pos_pos_trivial mod_neg_neg_trivial) qed spark_vc procedure_round_76 proof - let ?M = "4294967295 :: int" let ?INIT_CHAIN = "\h0 = ca_init, h1 = cb_init, h2 = cc_init, h3 = cd_init, h4 = ce_init\" have steps_to_steps': "steps (\n::nat. word_of_int (x (int n))) (from_chain ?INIT_CHAIN, from_chain ?INIT_CHAIN) 80 = from_chain_pair ( steps' (\left = ?INIT_CHAIN, right = ?INIT_CHAIN\) 80 x)" unfolding from_to_id by simp from \0 \ ca_init\ \ca_init \ ?M\ \0 \ cb_init\ \cb_init \ ?M\ \0 \ cc_init\ \cc_init \ ?M\ \0 \ cd_init\ \cd_init \ ?M\ \0 \ ce_init\ \ce_init \ ?M\ \0 \ cla\ \cla \ ?M\ \0 \ clb\ \clb \ ?M\ \0 \ clc\ \clc \ ?M\ \0 \ cld\ \cld \ ?M\ \0 \ cle\ \cle \ ?M\ \0 \ cra\ \cra \ ?M\ \0 \ crb\ \crb \ ?M\ \0 \ crc\ \crc \ ?M\ \0 \ crd\ \crd \ ?M\ \0 \ cre\ \cre \ ?M\ show ?thesis unfolding round_def unfolding steps_to_steps' unfolding H1[symmetric] by (simp add: uint_word_ariths(1) mod_simps - uint_word_of_int_id) + uint_word_of_int_id take_bit_int_eq_self) qed spark_end end diff --git a/src/HOL/Word/Conversions.thy b/src/HOL/Word/Conversions.thy deleted file mode 100644 --- a/src/HOL/Word/Conversions.thy +++ /dev/null @@ -1,556 +0,0 @@ -(* Author: Florian Haftmann, TUM -*) - -section \Proof of concept for conversions for algebraically founded bit word types\ - -theory Conversions - imports - Main - "HOL-Library.Type_Length" - "HOL-Library.Bit_Operations" - "HOL-Word.Word" -begin - -hide_const (open) unat uint sint word_of_int ucast scast - -context semiring_bits -begin - -lemma - exp_add_not_zero_imp_left: \2 ^ m \ 0\ - and exp_add_not_zero_imp_right: \2 ^ n \ 0\ if \2 ^ (m + n) \ 0\ -proof - - have \\ (2 ^ m = 0 \ 2 ^ n = 0)\ - proof (rule notI) - assume \2 ^ m = 0 \ 2 ^ n = 0\ - then have \2 ^ (m + n) = 0\ - by (rule disjE) (simp_all add: power_add) - with that show False .. - qed - then show \2 ^ m \ 0\ and \2 ^ n \ 0\ - by simp_all -qed - -lemma exp_not_zero_imp_exp_diff_not_zero: - \2 ^ (n - m) \ 0\ if \2 ^ n \ 0\ -proof (cases \m \ n\) - case True - moreover define q where \q = n - m\ - ultimately have \n = m + q\ - by simp - with that show ?thesis - by (simp add: exp_add_not_zero_imp_right) -next - case False - with that show ?thesis - by simp -qed - -end - - -subsection \Conversions to word\ - -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_eq_word_of_int [simp]: - \Word.Word = word_of_int\ - by (rule ext; transfer) simp - -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_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_of_nat_eq_0_iff [simp]: - \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ - using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) - -lemma word_of_int_eq_0_iff [simp]: - \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ - using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) - - -subsection \Conversion from word\ - -subsubsection \Generic unsigned conversion\ - -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 - -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) - -end - -context semiring_bits -begin - -lemma bit_unsigned_iff: - \bit (unsigned w) n \ 2 ^ n \ 0 \ 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 - -context semiring_bit_shifts -begin - -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 \2 ^ m \ 0\ - 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 \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ - by (metis (full_types) diff_add exp_add_not_zero_imp) - with True show ?thesis - by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff ac_simps exp_eq_zero_iff not_le) - next - case False - then show ?thesis - by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.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 Parity.bit_take_bit_iff) - -end - -context semiring_bit_operations -begin - -lemma unsigned_and_eq: - \unsigned (v AND w) = unsigned v AND unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) - -lemma unsigned_or_eq: - \unsigned (v OR w) = unsigned v OR unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) - -lemma unsigned_xor_eq: - \unsigned (v XOR w) = unsigned v XOR unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) - -end - -context ring_bit_operations -begin - -lemma unsigned_not_eq: - \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ - for w :: \'b::len word\ - by (rule bit_eqI) - (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) - -end - -lemma unsigned_of_nat [simp]: - \unsigned (word_of_nat n :: 'a::len word) = of_nat (take_bit LENGTH('a) n)\ - by transfer (simp add: nat_eq_iff take_bit_of_nat) - -lemma unsigned_of_int [simp]: - \unsigned (word_of_int n :: 'a::len word) = of_int (take_bit LENGTH('a) n)\ - by transfer (simp add: nat_eq_iff take_bit_of_nat) - -context unique_euclidean_semiring_numeral -begin - -lemma unsigned_greater_eq: - \0 \ unsigned w\ for w :: \'b::len word\ - by (transfer fixing: less_eq) simp - -lemma unsigned_less: - \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ - by (transfer fixing: less) (simp add: take_bit_int_less_exp) - -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_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)\) - (simp_all add: sbintrunc_minus_simps) - -lemma signed_minus_1 [simp]: - \signed (- 1 :: 'b::len word) = - 1\ - by (transfer fixing: uminus) 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) - -end - -context ring_bit_operations -begin - -lemma bit_signed_iff: - \bit (signed w) n \ 2 ^ n \ 0 \ 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) - 1) (push_bit n (signed w :: 'a))\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix m - assume \2 ^ m \ 0\ - define q where \q = LENGTH('b) - Suc 0\ - then have *: \LENGTH('b) = Suc q\ - by simp - show \bit (signed (push_bit n w)) m \ - bit (signed_take_bit (LENGTH('b) - 1) (push_bit n (signed w :: 'a))) m\ - proof (cases \q \ m\) - case True - moreover define r where \r = m - q\ - ultimately have \m = q + r\ - by simp - moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ - using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] - by simp_all - moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ - by (rule exp_not_zero_imp_exp_diff_not_zero) - ultimately show ?thesis - by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff - min_def * exp_eq_zero_iff le_diff_conv2) - next - case False - then show ?thesis - using exp_not_zero_imp_exp_diff_not_zero [of m n] - by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff - min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit - exp_eq_zero_iff) - qed -qed - -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) - -lemma signed_not_eq: - \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - define q where \q = LENGTH('b) - Suc 0\ - then have *: \LENGTH('b) = Suc q\ - by simp - show \bit (signed (NOT w)) n \ - bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ - proof (cases \q < n\) - case True - moreover define r where \r = n - Suc q\ - ultimately have \n = r + Suc q\ - by simp - moreover from \2 ^ n \ 0\ \n = r + Suc q\ - have \2 ^ Suc q \ 0\ - using exp_add_not_zero_imp_right by blast - ultimately show ?thesis - by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def - exp_eq_zero_iff) - next - case False - then show ?thesis - by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def - exp_eq_zero_iff) - qed -qed - -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 - -lemma signed_of_nat [simp]: - \signed (word_of_nat n :: 'a::len word) = of_int (signed_take_bit (LENGTH('a) - Suc 0) (int n))\ - by transfer simp - -lemma signed_of_int [simp]: - \signed (word_of_int n :: 'a::len word) = of_int (signed_take_bit (LENGTH('a) - Suc 0) n)\ - by transfer simp - - -subsubsection \Important special cases\ - -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 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 take_bit_int_less_exp) - 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 take_bit_int_less_exp) - 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 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 - -lemma of_int_sint [simp]: - \of_int (sint a) = signed a\ - by transfer (simp_all add: take_bit_signed_take_bit) - -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 take_bit_int_less_exp not_eq_complement mask_eq_exp_minus_1 OR_upper) - -context semiring_bit_shifts -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 Conversions.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_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\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ - by (simp add: min_def) - (metis (mono_tags) diff_diff_cancel local.exp_not_zero_imp_exp_diff_not_zero) - then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ - by (simp add: bit_signed_iff bit_unsigned_iff Conversions.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) -qed - -lemma signed_scast_eq: - \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ - by (simp add: min_def) - (metis (mono_tags) diff_diff_cancel local.exp_not_zero_imp_exp_diff_not_zero) - then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ - by (simp add: bit_signed_iff bit_unsigned_iff Conversions.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) -qed - -end - -end diff --git a/src/HOL/Word/More_Word.thy b/src/HOL/Word/More_Word.thy --- a/src/HOL/Word/More_Word.thy +++ b/src/HOL/Word/More_Word.thy @@ -1,50 +1,55 @@ (* Title: HOL/Word/More_thy *) section \Ancient comprehensive Word Library\ theory More_Word imports Word Ancient_Numeral Reversed_Bit_Lists Bits_Int Misc_Auxiliary Misc_Arithmetic Misc_set_bit Misc_lsb begin declare signed_take_bit_Suc [simp] lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq_decr_exp lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def +lemmas of_nat_word_eq_iff = word_of_nat_eq_iff +lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff +lemmas of_int_word_eq_iff = word_of_int_eq_iff +lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff + lemma shiftl_transfer [transfer_rule]: includes lifting_syntax shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" by (unfold shiftl_eq_push_bit) transfer_prover end diff --git a/src/HOL/Word/Word.thy b/src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy +++ b/src/HOL/Word/Word.thy @@ -1,4654 +1,5321 @@ (* Title: HOL/Word/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" Bits_Int Traditional_Syntax Bit_Comprehension Misc_Typedef 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 word_exp_length_eq_0 [simp]: \(2 :: 'a::len word) ^ LENGTH('a) = 0\ by transfer simp +lemma exp_eq_zero_iff: + \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ + by transfer simp + subsubsection \Basic code generation setup\ -lift_definition uint :: \'a::len word \ int\ +context +begin + +qualified lift_definition the_int :: \'a::len word \ int\ is \take_bit LENGTH('a)\ . +end + lemma [code abstype]: - \Word.Word (uint w) = w\ + \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 (uint v) (uint w)\ + \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ by transfer (simp add: equal) lemma [code]: - \uint 0 = 0\ + \Word.the_int 0 = 0\ by transfer simp lemma [code]: - \uint 1 = 1\ + \Word.the_int 1 = 1\ by transfer simp lemma [code]: - \uint (v + w) = take_bit LENGTH('a) (uint v + uint w)\ + \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]: - \uint (- w) = (let k = uint w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ + \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]: - \uint (v - w) = take_bit LENGTH('a) (uint v - uint w)\ + \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]: - \uint (v * w) = take_bit LENGTH('a) (uint v * uint w)\ + \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\ -lift_definition word_of_int :: \int \ 'a::len word\ - is \\k. k\ . - -lift_definition unat :: \'a::len word \ nat\ - is \nat \ take_bit LENGTH('a)\ +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 [simp]: + \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 [simp]: + \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 -lift_definition sint :: \'a::len word \ int\ - \ \treats the most-significant bit as a sign bit\ - is \signed_take_bit (LENGTH('a) - 1)\ - by (simp add: signed_take_bit_decr_length_iff) +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 [simp]: + \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 [simp]: + \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) + +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)\) + (simp_all add: sbintrunc_minus_simps) + +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 [simp]: + \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 [simp]: + \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) + +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 of_nat_word_eq_iff: - \of_nat m = (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ +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 of_nat_word_eq_0_iff: - \of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ - using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) - -lemma of_int_word_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_eq_0_iff: - \of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ - using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) +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 + +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 \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) - 1) \ P (2 * a)\ - and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - 1) \ P (1 + 2 * a)\ + 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) - 1\ + 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 (unfold unat_def) (transfer, simp add: take_bit_eq_mod) + 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 of_nat_word_eq_0_iff l) - moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\ + by (auto simp add: word_greater_zero_iff l) + 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) - 1)\ + 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) 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) 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 show \a mod b div b = 0\ for a b :: \'a word\ apply transfer apply (simp add: take_bit_eq_mod) apply (subst (3) mod_pos_pos_trivial [of _ \2 ^ LENGTH('a)\]) apply simp_all apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) using pos_mod_bound [of \2 ^ LENGTH('a)\] apply simp proof - fix aa :: int and ba :: int have f1: "\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 "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \ ba mod 2 ^ len_of (TYPE('a)::'a itself) \ 0 \ aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) qed 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: mod_2_eq_odd 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 (auto simp add: take_bit_of_mask even_mask_div_iff) 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 + instantiation word :: (len) semiring_bit_shifts begin 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 proof show \push_bit n a = a * 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp add: push_bit_eq_mult) show \drop_bit n a = a div 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) show \take_bit n a = a mod 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (auto simp flip: take_bit_eq_mod) 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\ +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) + + +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 + . + +instance by (standard; transfer) + (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 + bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) + +end + +lemma [code_abbrev]: + \push_bit n 1 = (2 :: 'a::len word) ^ n\ + by (fact push_bit_of_1) + +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 + +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 + subsection \Conversions including casts\ +subsubsection \Generic unsigned conversion\ + +context semiring_bits +begin + +lemma bit_unsigned_iff: + \bit (unsigned w) n \ 2 ^ n \ 0 \ 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 + +context semiring_bit_shifts +begin + +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 \2 ^ m \ 0\ + 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 \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ + by (metis (full_types) diff_add exp_add_not_zero_imp) + with True show ?thesis + by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps) + next + case False + then show ?thesis + by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.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 Parity.bit_take_bit_iff) + +end + +context semiring_bit_operations +begin + +lemma unsigned_and_eq: + \unsigned (v AND w) = unsigned v AND unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) + +lemma unsigned_or_eq: + \unsigned (v OR w) = unsigned v OR unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) + +lemma unsigned_xor_eq: + \unsigned (v XOR w) = unsigned v XOR unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) + +end + +context ring_bit_operations +begin + +lemma unsigned_not_eq: + \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ + for w :: \'b::len word\ + by (rule bit_eqI) + (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) + +end + +context unique_euclidean_semiring_numeral +begin + +lemma unsigned_greater_eq: + \0 \ unsigned w\ for w :: \'b::len word\ + by (transfer fixing: less_eq) simp + +lemma unsigned_less: + \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 (signed w) n \ 2 ^ n \ 0 \ 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\ +proof (rule bit_eqI) + fix m + assume \2 ^ m \ 0\ + define q where \q = LENGTH('b) - Suc 0\ + then have *: \LENGTH('b) = Suc q\ + by simp + show \bit (signed (push_bit n w)) m \ + bit (signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))) m\ + proof (cases \q \ m\) + case True + moreover define r where \r = m - q\ + ultimately have \m = q + r\ + by simp + moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ + using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] + by simp_all + moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ + by (rule exp_not_zero_imp_exp_diff_not_zero) + ultimately show ?thesis + by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff + min_def * exp_eq_zero_iff le_diff_conv2) + next + case False + then show ?thesis + using exp_not_zero_imp_exp_diff_not_zero [of m n] + by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff + min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit + exp_eq_zero_iff) + qed +qed + +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) + +lemma signed_not_eq: + \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + define q where \q = LENGTH('b) - Suc 0\ + then have *: \LENGTH('b) = Suc q\ + by simp + show \bit (signed (NOT w)) n \ + bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ + proof (cases \q < n\) + case True + moreover define r where \r = n - Suc q\ + ultimately have \n = r + Suc q\ + by simp + moreover from \2 ^ n \ 0\ \n = r + Suc q\ + have \2 ^ Suc q \ 0\ + using exp_add_not_zero_imp_right by blast + ultimately show ?thesis + by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def + exp_eq_zero_iff) + next + case False + then show ?thesis + by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def + exp_eq_zero_iff) + qed +qed + +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 + + +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 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_shifts +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 exp_eq_zero_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\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ + by (simp add: min_def) + (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) + then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ + by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) +qed + +lemma signed_scast_eq: + \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ + by (simp add: min_def) + (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) + then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ + by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) +qed + +end + lemma uint_nonnegative: "0 \ uint w" - by transfer simp + by (fact unsigned_greater_eq) lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len word" - by transfer (simp add: take_bit_eq_mod) + by (fact unsigned_less) lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len word" - using uint_nonnegative uint_bounded by (rule mod_pos_pos_trivial) + by transfer (simp add: take_bit_eq_mod) lemma word_uint_eqI: "uint a = uint b \ a = b" - by transfer simp + by (fact unsigned_word_eqI) lemma word_uint_eq_iff: "a = b \ uint a = uint b" - using word_uint_eqI by auto - -lemma Word_eq_word_of_int [code_post]: - \Word.Word = word_of_int\ - by rule (transfer, rule) - -lemma uint_word_of_int_eq [code]: + 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 add: word_of_int_uint) + then show "PROP P x" + by (simp only: word_of_int_uint) qed -lemma sint_uint [code]: - \sint w = signed_take_bit (LENGTH('a) - 1) (uint w)\ +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 [code]: +lemma unat_eq_nat_uint: \unat w = nat (uint w)\ by simp -lift_definition ucast :: \'a::len word \ 'b::len word\ - is \take_bit LENGTH('a)\ - by simp - -lemma ucast_eq [code]: +lemma ucast_eq: \ucast w = word_of_int (uint w)\ by transfer simp -lift_definition scast :: \'a::len word \ 'b::len word\ - is \signed_take_bit (LENGTH('a) - 1)\ - by (simp flip: signed_take_bit_decr_length_iff) - -lemma scast_eq [code]: +lemma scast_eq: \scast w = word_of_int (sint w)\ by transfer simp -lemma uint_0_eq [simp]: +lemma uint_0_eq: \uint 0 = 0\ - by transfer simp - -lemma uint_1_eq [simp]: + by (fact unsigned_0) + +lemma uint_1_eq: \uint 1 = 1\ - by transfer simp + by (fact unsigned_1) lemma word_m1_wi: "- 1 = word_of_int (- 1)" - by transfer rule + by simp lemma uint_0_iff: "uint x = 0 \ x = 0" - by (simp add: word_uint_eq_iff) + by (auto simp add: unsigned_word_eqI) lemma unat_0_iff: "unat x = 0 \ x = 0" - by transfer (auto intro: antisym) - -lemma unat_0 [simp]: "unat 0 = 0" - by transfer simp + 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 [simp]: "ucast 0 = 0" - by transfer simp - -lemma sint_0 [simp]: "sint 0 = 0" - by (simp add: sint_uint) - -lemma scast_0 [simp]: "scast 0 = 0" - by transfer simp - -lemma sint_n1 [simp] : "sint (- 1) = - 1" - by transfer simp - -lemma scast_n1 [simp]: "scast (- 1) = - 1" - by transfer simp +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 [simp]: "unat (1::'a::len word) = 1" - by transfer simp - -lemma ucast_1 [simp]: "ucast (1::'a::len word) = 1" - by transfer simp +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\ 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 word_of_nat: "of_nat n = word_of_int (int n)" - by (induct n) (auto simp add : word_of_int_hom_syms) - -lemma word_of_int: "of_int = word_of_int" - apply (rule ext) - apply (case_tac x rule: int_diff_cases) - apply (simp add: word_of_nat wi_hom_sub) - done - -lemma word_of_int_eq: - "word_of_int = of_int" - by (rule ext) (transfer, rule) - -definition udvd :: "'a::len word \ 'a::len word \ bool" (infixl "udvd" 50) - where "a udvd b = (\n\0. uint b = n * uint a)" - -lemma exp_eq_zero_iff: - \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ - by transfer simp - 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\ (\(_/ <=s _)\ [50, 51] 50) is \\k l. signed_take_bit (LENGTH('a) - 1) k \ signed_take_bit (LENGTH('a) - 1) l\ by (simp flip: signed_take_bit_decr_length_iff) lemma word_sle_eq [code]: \a <=s b \ sint a \ sint b\ by transfer simp lift_definition word_sless :: \'a::len word \ 'a word \ bool\ (\(_/ [50, 51] 50) is \\k l. signed_take_bit (LENGTH('a) - 1) k < signed_take_bit (LENGTH('a) - 1) l\ by (simp flip: signed_take_bit_decr_length_iff) lemma word_sless_eq: \x x <=s y \ x \ y\ by transfer (simp add: signed_take_bit_decr_length_iff less_le) lemma [code]: \a sint a < sint b\ by transfer simp lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) 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 (rule signed_linorder) lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len word" by transfer simp lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 bintr_lt2p) 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\ - -lemma uint_take_bit_eq [code]: +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) lift_definition shiftl1 :: \'a::len word \ 'a word\ is \(*) 2\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by transfer (simp add: take_bit_eq_mod mod_simps) lemma shiftl1_eq_mult_2: \shiftl1 = (*) 2\ by (rule ext, transfer) simp lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) lift_definition shiftr1 :: \'a::len word \ 'a word\ \ \shift right as unsigned or as signed, ie logical or arithmetic\ is \\k. take_bit LENGTH('a) k div 2\ by simp lemma shiftr1_eq_div_2: \shiftr1 w = w div 2\ by transfer simp lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by transfer (auto simp flip: bit_Suc simp add: bit_take_bit_iff) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ 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 - . - -instance by (standard; transfer) - (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 - bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) - -end - -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 - instantiation word :: (len) semiring_bit_syntax begin lift_definition test_bit_word :: \'a::len 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 lemma test_bit_word_eq: \test_bit = (bit :: 'a word \ _)\ by transfer simp 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 [code]: \test_bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: test_bit_word_eq bit_word_iff_drop_bit_and) lift_definition shiftl_word :: \'a::len word \ nat \ 'a word\ is \\k n. push_bit n k\ 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 lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by transfer rule lift_definition shiftr_word :: \'a::len word \ nat \ 'a word\ is \\k n. drop_bit n (take_bit LENGTH('a) k)\ by simp lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by transfer simp instance by (standard; transfer) simp_all end lemma shiftl_code [code]: \w << n = w * 2 ^ n\ for w :: \'a::len word\ by transfer (simp add: push_bit_eq_mult) lemma shiftl1_code [code]: \shiftl1 w = w << 1\ by transfer (simp add: push_bit_eq_mult ac_simps) -lemma uint_shiftr_eq [code]: +lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ - for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) +lemma [code]: + \Word.the_int (w >> n) = drop_bit n (Word.the_int w)\ + by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) + lemma shiftr1_code [code]: \shiftr1 w = w >> 1\ by transfer (simp add: drop_bit_Suc) lemma word_test_bit_def: \test_bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemma shiftl_def: \w << n = (shiftl1 ^^ n) w\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = (shiftr1 ^^ n) w\ proof - have \drop_bit n = (((\k::int. k div 2) ^^ n))\ for n by (rule sym, induction n) (simp_all add: fun_eq_iff drop_bit_Suc flip: drop_bit_half) then show ?thesis apply transfer apply simp apply (metis bintrunc_bintrunc rco_bintr) done qed lemma bit_shiftl_word_iff: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: shiftl_word_eq bit_push_bit_iff exp_eq_zero_iff not_le) lemma [code]: \push_bit n w = w << n\ for w :: \'a::len word\ by (simp add: shiftl_word_eq) +lemma [code]: + \drop_bit n w = w >> n\ for w :: \'a::len word\ + by (simp add: shiftr_word_eq) + lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: shiftr_word_eq bit_drop_bit_eq) -lemma [code]: - \drop_bit n w = w >> n\ for w :: \'a::len word\ - by (simp add: shiftr_word_eq) - -lemma [code]: - \uint (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (uint w) else uint w)\ - for w :: \'a::len word\ - by transfer (simp add: min_def) - -lemma [code]: - \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ - by transfer simp - -lemma [code_abbrev]: - \push_bit n 1 = (2 :: 'a::len word) ^ n\ - by (fact push_bit_of_1) - lemma - word_not_def [code]: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" + 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)+ -lemma [code abstract]: - \uint (v AND w) = uint v AND uint w\ - by transfer simp - -lemma [code abstract]: - \uint (v OR w) = uint v OR uint w\ - by transfer simp - -lemma [code abstract]: - \uint (v XOR w) = uint v XOR uint w\ - by transfer simp - lift_definition setBit :: \'a::len word \ nat \ 'a word\ is \\k n. set_bit n k\ by (simp add: take_bit_set_bit_eq) lemma set_Bit_eq: \setBit w n = set_bit n w\ by transfer simp lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by transfer (auto simp add: bit_set_bit_iff) lift_definition clearBit :: \'a::len word \ nat \ 'a word\ is \\k n. unset_bit n k\ by (simp add: take_bit_unset_bit_eq) lemma clear_Bit_eq: \clearBit w n = unset_bit n w\ by transfer simp lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_unset_bit_iff) 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 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 \Type-definition locale instantiations\ definition uints :: "nat \ int set" \ \the sets of integers representing the words\ where "uints n = range (take_bit n)" definition sints :: "nat \ int set" where "sints n = range (signed_take_bit (n - 1))" lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" by (simp add: uints_def range_bintrunc) lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" by (simp add: sints_def range_sbintrunc) definition unats :: "nat \ nat set" where "unats n = {i. i < 2 ^ n}" \ \naturals\ lemma uints_unats: "uints n = int ` unats n" apply (unfold unats_def uints_num) apply safe apply (rule_tac image_eqI) apply (erule_tac nat_0_le [symmetric]) by auto lemma unats_uints: "unats n = nat ` uints n" by (auto simp: uints_unats image_iff) lemma td_ext_uint: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (\w::int. w mod 2 ^ LENGTH('a))" apply (unfold td_ext_def') apply transfer apply (simp add: uints_num take_bit_eq_mod) done interpretation word_uint: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "\w. w mod 2 ^ LENGTH('a::len)" by (fact td_ext_uint) lemmas td_uint = word_uint.td_thm lemmas int_word_uint = word_uint.eq_norm lemma td_ext_ubin: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (take_bit (LENGTH('a)))" apply standard apply transfer apply simp done interpretation word_ubin: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "take_bit (LENGTH('a::len))" by (fact td_ext_ubin) lemma td_ext_unat [OF refl]: "n = LENGTH('a::len) \ td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" apply (standard; transfer) - apply (simp_all add: unats_def take_bit_int_less_exp take_bit_of_nat take_bit_nat_eq_self_iff + apply (simp_all add: unats_def take_bit_of_nat take_bit_nat_eq_self_iff flip: take_bit_eq_mod) done lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] interpretation word_unat: td_ext "unat::'a::len word \ nat" of_nat "unats (LENGTH('a::len))" "\i. i mod 2 ^ LENGTH('a::len)" by (rule td_ext_unat) lemmas td_unat = word_unat.td_thm lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" for z :: "'a::len word" apply (unfold unats_def) apply clarsimp apply (rule xtrans, rule unat_lt2p, assumption) done lemma td_ext_sbin: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (signed_take_bit (LENGTH('a) - 1))" - apply (unfold td_ext_def' sint_uint) - apply (simp add : word_ubin.eq_norm) - apply (cases "LENGTH('a)") - apply (auto simp add : sints_def) - apply (rule sym [THEN trans]) - apply (rule word_ubin.Abs_norm) - apply (simp only: bintrunc_sbintrunc) - apply (drule sym) - apply simp - done + by (standard; transfer) (auto simp add: sints_def) lemma td_ext_sint: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1))" using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) text \ We do \sint\ before \sbin\, before \sint\ is the user version and interpretations do not produce thm duplicates. I.e. we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, because the latter is the same thm as the former. \ interpretation word_sint: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - 2 ^ (LENGTH('a::len) - 1)" by (rule td_ext_sint) interpretation word_sbin: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "signed_take_bit (LENGTH('a::len) - 1)" by (rule td_ext_sbin) lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] lemmas td_sint = word_sint.td subsection \More shift operations\ lift_definition sshiftr1 :: \'a::len word \ 'a word\ is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) k div 2)\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - 1) k))\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ by (fact arg_cong) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by transfer simp lemma sshiftr_eq: \w >>> n = (sshiftr1 ^^ n) w\ proof - have *: \(\k::int. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)) ^^ Suc n = take_bit LENGTH('a) \ drop_bit (Suc n) \ signed_take_bit (LENGTH('a) - Suc 0)\ for n apply (induction n) apply (auto simp add: fun_eq_iff drop_bit_Suc) apply (metis (no_types, lifting) Suc_pred funpow_swap1 len_gt_0 sbintrunc_bintrunc sbintrunc_rest) done show ?thesis apply transfer apply simp subgoal for k n apply (cases n) apply (simp_all only: *) apply simp_all done done qed lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) -lemma uint_sshiftr_eq [code]: +lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) +lemma [code]: + \Word.the_int (w >>> n) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ + for w :: \'a::len word\ + by transfer simp + lemma sshift1_code [code]: \sshiftr1 w = w >>> 1\ by transfer (simp add: drop_bit_Suc) 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 apply (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 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 apply (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 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 apply (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 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 (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 (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 (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 [code]: +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\ apply transfer apply (simp add: concat_bit_def take_bit_drop_bit push_bit_take_bit min_def) using mod_less_divisor not_less apply blast done +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 lemma bit_word_cat_iff: \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 a = (case bin_split (LENGTH('c)) (uint a) of (u, v) \ (word_of_int u, word_of_int v))" definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done definition word_rsplit :: "'a::len word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" abbreviation (input) max_word :: \'a::len word\ \ \Largest representable machine integer.\ where "max_word \ - 1" subsection \Theorems about typedefs\ lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by (auto simp: sint_uint word_ubin.eq_norm sbintrunc_bintrunc_lt) lemma uint_sint: "uint w = take_bit (LENGTH('a)) (sint w)" for w :: "'a::len word" by (auto simp: sint_uint bintrunc_sbintrunc_le) lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: bintrunc_bintrunc_min word_size) apply (simp add: min.absorb2) done lemma wi_bintr: "LENGTH('a::len) \ n \ word_of_int (take_bit n w) = (word_of_int w :: 'a word)" by (auto simp: word_ubin.norm_eq_iff [symmetric] min.absorb1) lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" by (fact uints_def [unfolded no_bintr_alt1]) 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)" unfolding word_numeral_alt by (rule word_ubin.eq_norm) lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" by (simp only: word_neg_numeral_alt word_ubin.eq_norm) lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" by (simp only: word_numeral_alt word_sbin.eq_norm) lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" by (simp only: word_neg_numeral_alt word_sbin.eq_norm) 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" apply (unfold word_size) apply (rule word_uint.Rep_eqD) apply (rule box_equals) defer apply (rule word_ubin.norm_Rep)+ apply simp done lemma uint_ge_0 [iff]: "0 \ uint x" for x :: "'a::len word" using word_uint.Rep [of x] by (simp add: uints_num) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len word" using word_uint.Rep [of x] by (simp add: uints_num) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) 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" by (metis bintr_lt2p bintr_uint) 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 only: word_numeral_alt int_word_uint) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" by (simp only: word_neg_numeral_alt int_word_uint) 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)" unfolding word_numeral_alt by (rule int_word_sint) lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" unfolding word_0_wi .. lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" unfolding word_1_wi .. 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 (simp only: word_numeral_alt) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" by (simp only: word_numeral_alt wi_hom_syms) 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) lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" unfolding word_size by (rule uint_range') lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" unfolding word_size by (rule sint_range') 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]) subsection \Testing bits\ lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len word" unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) lemma test_bit_size [rule_format] : "w !! n \ n < size w" for w :: "'a::len word" apply (unfold word_test_bit_def) apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: nth_bintr word_size) apply fast done lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) for x y :: "'a::len word" -proof - assume ?P - then show ?Q - by simp -next - assume ?Q - then have *: \bit (uint x) n \ bit (uint y) n\ if \n < LENGTH('a)\ for n - using that by (simp add: word_test_bit_def) - show ?P - proof (rule word_uint_eqI, rule bit_eqI, rule iffI) - fix n - assume \bit (uint x) n\ - then have \n < LENGTH('a)\ - by (simp add: bit_take_bit_iff uint.rep_eq) - with * \bit (uint x) n\ - show \bit (uint y) n\ - by simp - next - fix n - assume \bit (uint y) n\ - then have \n < LENGTH('a)\ - by (simp add: bit_take_bit_iff uint.rep_eq) - with * \bit (uint y) n\ - show \bit (uint x) n\ - by simp - qed -qed + by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) lemma word_eqI: "(\n. n < size u \ u !! n = v !! n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ u !! x = v !! x" for u v :: "'a::len word" by simp lemma test_bit_bin': "w !! n \ n < size w \ bin_nth (uint w) n" by (simp add: word_test_bit_def word_size nth_bintr [symmetric]) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma bin_nth_uint_imp: "bin_nth (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) apply (subst word_ubin.norm_Rep) apply assumption done lemma bin_nth_sint: "LENGTH('a) \ n \ bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" apply (subst word_sbin.norm_Rep [symmetric]) apply (auto simp add: nth_sbintr) done lemmas bintr_num = word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas sbintr_num = word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemma num_of_bintr': "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding bintr_num by (erule subst, simp) lemma num_of_sbintr': "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding sbintr_num by (erule subst, simp) lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (take_bit (LENGTH('a::len)) (numeral x))" by (simp only: word_ubin.Abs_norm word_numeral_alt) lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" by (simp only: word_sbin.Abs_norm word_numeral_alt) 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) \ Parity.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 lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) 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 exp_eq_zero_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 test_bit_def' = word_test_bit_def [THEN fun_cong] 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" apply (cases \LENGTH('a)\) apply simp_all apply (simp add: bit_iff_odd_drop_bit) apply transfer apply (simp add: take_bit_drop_bit) apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) apply (auto elim!: evenE) apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) done subsection \Word Arithmetic\ -lemma udvdI: "0 \ n \ uint b = n * uint a \ a udvd b" - by (auto simp: udvd_def) - 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 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)" - apply (simp_all only: word_arith_wis) - apply (simp_all add: word_uint.eq_norm) - done + 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" apply (simp_all only: word_sbin.inverse_norm [symmetric]) apply (simp_all add: wi_hom_syms) apply transfer apply simp apply transfer apply simp done lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] 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\ -lemma udvd_nat_alt: "a udvd b \ (\n\0. unat b = n * unat a)" - supply nat_uint_eq [simp del] - apply (unfold udvd_def) - apply safe - apply (simp add: unat_eq_nat_uint nat_mult_distrib) - apply (simp add: uint_nat) - apply (rule exI) - apply safe - prefer 2 - apply (erule notE) - apply (rule refl) - apply force +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\ unat_lt2p [of w] + show False + by simp + qed + ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ + by (auto simp add: take_bit_nat_eq_self_iff 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_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)\ + apply (auto elim!: dvdE simp add: udvd_iff_dvd) + apply (simp only: uint_nat) + apply auto + using of_nat_0_le_iff apply blast + apply (simp only: unat_eq_nat_uint) + apply (simp add: nat_mult_distrib) done -lemma udvd_iff_dvd: "x udvd y \ unat x dvd unat y" - unfolding dvd_def udvd_nat_alt by force - 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 int_word_uint word_arith_wis mod_diff_right_eq) 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 (mono_tags, hide_lams) diff_ge_0_iff_ge mod_pos_pos_trivial of_nat_0_le_iff take_bit_eq_mod uint_nat uint_sub_lt2p word_sub_wi word_ubin.eq_norm) find_theorems uint \- _\ 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 add: int_mod_ge) 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 (auto simp add: not_less) apply (rule antisym) apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) apply (rule int_mod_ge) apply simp_all done 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 apply (auto simp add: not_le) apply (rule antisym) apply (simp only: flip: mod_add_self2 [of \x - y\ z]) apply (rule zmod_le_nonneg_dividend) apply simp apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) done 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) subsection \Definition of \uint_arith\\ lemma word_of_int_inverse: "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" for a :: "'a::len word" apply (erule word_uint.Abs_inverse' [rotated]) apply (simp add: uints_num) done 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 take_bit_int_less_exp) + 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 dest!: word_of_int_inverse - simp: int_word_uint - split: uint_split) + by auto (metis take_bit_int_eq_self_iff) lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint.Rep_inject [symmetric] uint_sub_if' uint_plus_if' \ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ lemma power_False_cong: "False \ a ^ b = c ^ d" by auto \ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ ML \ fun uint_arith_simpset ctxt = ctxt addsimps @{thms uint_arith_simps} delsimps @{thms word_uint.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} 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 (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" unfolding word_size by uint_arith 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 uint_arith 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 uint_arith lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" for x :: "'a::len word" by uint_arith lemma sub_wrap_lt: "x < x - z \ x < z" for x z :: "'a::len word" by uint_arith lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" for x z :: "'a::len word" by uint_arith 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" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (erule uint_add_le [THEN order_trans]) done lemma udvd_decr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p \ q - K" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (drule le_diff_eq [THEN iffD2]) apply (erule order_trans) apply (rule uint_sub_ge) done 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" - apply (unfold udvd_def) + apply (unfold udvd_unfold_int) apply clarify apply (erule (2) udvd_decr0) done 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" supply [[simproc del: linordered_ring_less_cancel_factor]] - apply (unfold udvd_def) + apply (unfold udvd_unfold_int) apply clarify apply (simp add: uint_arith_simps split: if_split_asm) prefer 2 apply (insert uint_range' [of s])[1] apply arith apply (drule add.commute [THEN xtrans(1)]) apply (simp flip: diff_less_eq) apply (subst (asm) mult_less_cancel_right) apply simp apply (simp add: diff_eq_eq not_less) apply (subst (asm) (3) zless_iff_Suc_zadd) apply auto apply (auto simp add: algebra_simps) apply (drule less_le_trans [of _ \2 ^ LENGTH('a)\]) apply assumption apply (simp add: mult_less_0_iff) done 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 add: word_of_int) + 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))" using word_ubin.norm_eq_iff [where 'a='a, of "numeral bin" 0] by (simp add: iszero_def [symmetric]) 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)" apply (rule allI) apply (rule word_unat.Abs_cases) apply (unfold unats_def) apply auto done 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 add: word_unat.inverse_norm) 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: word_of_nat wi_hom_mult) + by (simp add: wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" - by (simp add: word_of_nat wi_hom_succ ac_simps) + 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 (simp add: word_div_def word_of_nat zdiv_int uint_nat) - + 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 (simp add: word_mod_def word_of_nat zmod_int uint_nat) + 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 simp 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" apply (auto simp: unat_word_ariths) apply (metis unat_lt2p word_unat.eq_norm) done lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" apply (auto simp: unat_word_ariths) apply (metis unat_lt2p word_unat.eq_norm) done 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: unat_word_ariths not_less) apply (subst (asm) le_iff_add) apply auto apply (metis add_less_cancel_left add_less_cancel_right le_less_trans less_imp_le mod_less unat_lt2p) done lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len word" apply (erule order_trans) apply (erule olen_add_eqv [THEN iffD1]) done lemmas un_ui_le = trans [OF word_le_nat_alt [symmetric] word_le_def] 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)" supply nat_uint_eq [simp del] apply (unfold word_size) apply (simp add: un_ui_le) apply (auto simp add: unat_eq_nat_uint uint_sub_if') apply (rule nat_diff_distrib) prefer 3 apply (simp add: algebra_simps) apply (rule nat_diff_distrib [THEN trans]) prefer 3 apply (subst nat_add_distrib) prefer 3 apply (simp add: nat_power_eq) apply auto apply uint_arith done lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma uint_div: \uint (x div y) = uint x div uint y\ - by (metis div_le_dividend le_less_trans mod_less uint_nat unat_lt2p unat_word_ariths(6) zdiv_int) + by (fact uint_div_distrib) lemma unat_div: \unat (x div y) = unat x div unat y\ - by (simp add: uint_div nat_div_distrib flip: nat_uint_eq) + by (fact unat_div_distrib) lemma uint_mod: \uint (x mod y) = uint x mod uint y\ - by (metis (no_types, lifting) le_less_trans mod_by_0 mod_le_divisor mod_less neq0_conv uint_nat unat_lt2p unat_word_ariths(7) zmod_int) + by (fact uint_mod_distrib) lemma unat_mod: \unat (x mod y) = unat x mod unat y\ - by (simp add: uint_mod nat_mod_distrib flip: nat_uint_eq) + by (fact unat_mod_distrib) text \Definition of \unat_arith\ tactic\ 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: unat_of_nat) + by auto (metis take_bit_nat_eq_self_iff) 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: unat_of_nat) + by auto (metis take_bit_nat_eq_self_iff) lemmas of_nat_inverse = word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat.Rep_inject [symmetric] unat_sub_if' unat_plus_if' unat_div unat_mod \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ fun unat_arith_simpset ctxt = ctxt addsimps @{thms unat_arith_simps} delsimps @{thms word_unat.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} 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 (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" 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" apply unat_arith apply clarsimp apply (subst unat_mult_lem [THEN iffD1]) apply auto done lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply unat_arith apply clarsimp apply (drule mult_le_mono1) apply (erule order_le_less_trans) apply (metis add_lessD1 div_mult_mod_eq unat_lt2p) done 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" apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule (1) mult_less_mono1) apply (erule order_less_le_trans) apply auto done lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" for i k x :: "'a::len word" apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule mult_le_mono1) apply (erule order_trans) apply auto done lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply (unfold uint_nat) apply (drule div_lt') apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) done 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_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) 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 lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] 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)}\ apply (rule inj_onI) apply transfer apply (simp add: take_bit_eq_mod) done lemma inj_uint: \inj uint\ by (rule injI) simp lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ by transfer (auto simp add: bintr_lt2p range_bintrunc) lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ proof - have \uint ` (UNIV :: 'a word set) = uint ` (word_of_int :: int \ 'a word) ` {0..<2 ^ LENGTH('a::len)}\ by transfer (simp add: range_bintrunc, auto simp add: take_bit_eq_mod) then show ?thesis using inj_image_eq_iff [of \uint :: 'a word \ int\ \UNIV :: 'a word set\ \word_of_int ` {0..<2 ^ LENGTH('a)} :: 'a word set\, OF inj_uint] by simp qed 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) instance word :: (len) finite by standard (simp add: UNIV_eq) subsection \Bitwise Operations on Words\ lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def 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)" by (transfer, simp)+ 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 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 lemma test_bit_wi [simp]: "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth x n" by (simp add: word_test_bit_def word_ubin.eq_norm nth_bintr) lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (test_bit :: 'a::len word \ _)" by (simp only: test_bit_eq_bit) transfer_prover lemma word_ops_nth_size: "n < size x \ (x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n) \ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ x !! n)" for x :: "'a::len word" unfolding word_size by transfer (simp add: bin_nth_ops) lemma word_ao_nth: "(x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n)" for x :: "'a::len word" by transfer (auto simp add: bin_nth_ops) lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_numeral [simp]: "(numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth (numeral w) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth (- numeral w) n" by transfer (rule refl) lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" by transfer auto lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" by transfer simp lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" 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 (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) 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 (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) 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 (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) 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 simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" by simp lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" by transfer (simp add: bin_add_not) 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 (auto simp: word_le_def uint_or intro: le_int_or) 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 (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 (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 lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: nth_sbintr word_test_bit_def [symmetric]) lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" by transfer (simp add: bin_sc_eq) lemma clearBit_no [simp]: "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" by transfer (simp add: bin_sc_eq) lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" - by (auto simp: word_test_bit_def word_ubin.eq_norm nth_bintr nth_2p_bin) + by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" - by (simp add: test_bit_2p [symmetric] word_of_int [symmetric]) + by transfer (auto simp add: bit_exp_iff) lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" - apply (unfold word_arith_power_alt) - apply (case_tac "LENGTH('a)") - apply clarsimp - apply (case_tac "nat") - apply clarsimp - apply (case_tac "n") - apply clarsimp - apply clarsimp - apply (drule word_gt_0 [THEN iffD1]) - apply (safe intro!: word_eqI) - apply (auto simp add: nth_2p_bin) - apply (erule notE) - apply (simp (no_asm_use) add: uint_word_of_int word_size) - apply (subst mod_pos_pos_trivial) - apply simp - apply (rule power_strict_increasing) - apply simp_all + apply (cases \n < LENGTH('a)\; transfer) + apply auto done lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) lemma bang_is_le: "x !! m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done subsection \Bit comprehension\ instantiation word :: (len) bit_comprehension begin definition word_set_bits_def: \(BITS n. P n) = (horner_sum of_bool 2 (map P [0.. instance .. end lemma bit_set_bits_word_iff: \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) lemma set_bits_bit_eq: \set_bits (bit w) = w\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_set_bits_word_iff bit_imp_le_length) lemma set_bits_K_False [simp]: \set_bits (\_. False) = (0 :: 'a :: len word)\ by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) interpretation test_bit: td_ext "(!!) :: 'a::len word \ nat \ bool" set_bits "{f. \i. f i \ i < LENGTH('a::len)}" "(\h i. h i \ i < LENGTH('a::len))" by standard (auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) lemmas td_nth = test_bit.td_thm subsection \Shifting, Rotating, and Splitting Words\ lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0 [simp] : "shiftl1 0 = 0" by transfer simp lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by (simp add: shiftl1_def_u wi_hom_syms) lemma shiftr1_0 [simp]: "shiftr1 0 = 0" by transfer simp lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" by transfer simp lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" by transfer simp lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" by transfer simp lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" by transfer simp lemma sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" by transfer simp lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" by transfer (auto simp add: bit_double_iff) lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct "m" arbitrary: n) apply (auto simp add: nth_shiftr1) done text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), where \f\ (ie \bin_rest\) takes normal arguments to normal results, thus we get (2) from (1) \ lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" by transfer simp lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma bit_sshiftr_word_iff: \bit (w >>> m) 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_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma nth_sshiftr : "sshiftr w m !! n = (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by transfer simp lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" apply (unfold shiftr_def) apply (induct n) apply simp apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" apply transfer apply (rule bit_eqI) apply (simp add: bit_signed_take_bit_iff bit_drop_bit_eq min_def flip: drop_bit_eq_div) done lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ apply transfer apply (simp add: bit_take_bit_iff flip: bit_Suc) apply (subst disjunctive_add) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff simp flip: bit_Suc) done subsubsection \shift functions in terms of lists of bools\ lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" apply (rule bit_word_eqI) apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) done lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemma shiftl_numeral [simp]: \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ by (fact shiftl_word_eq) lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done \ \note -- the following results use \'a::len word < number_ring\\ lemma shiftl1_2t: "shiftl1 w = 2 * w" for w :: "'a::len word" by (simp add: shiftl1_eq wi_hom_mult [symmetric]) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (simp add: shiftl1_2t) lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (take_bit (LENGTH('a)) (numeral w)))" - unfolding shiftr1_eq word_numeral_alt by (simp add: word_ubin.eq_norm) + by transfer simp lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (signed_take_bit (LENGTH('a) - 1) (numeral w)))" unfolding sshiftr1_eq word_numeral_alt by (simp add: word_sbin.eq_norm) text \TODO: rules for \<^term>\- (numeral n)\\ 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 shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) lemma sshiftr_numeral [simp]: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ apply (rule word_eqI) apply (cases \LENGTH('a)\) apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr nth_sbintr not_le not_less less_Suc_eq_le ac_simps) done lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" apply (induct ys arbitrary: n) apply simp_all apply (case_tac n) apply simp_all done lemma align_lem_or [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (|) x y = take m x @ drop m y" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def) done lemma align_lem_and [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (\) x y = replicate (n + m) False" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def map_replicate_const) done 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 exp_eq_zero_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 (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ by (simp add: bit_mask_iff exp_eq_zero_iff not_le) end lemma nth_mask [simp]: \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))" by (auto simp add: nth_bintr word_size intro: word_eqI) lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" apply (rule word_eqI) apply (simp add: nth_bintr word_size word_ops_nth_size) apply (auto simp add: test_bit_bin) done lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) 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: nth_bintr word_size word_ops_nth_size word_eq_iff) 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" apply (simp add: uint_and uint_mask_eq) apply (rule AND_upper2'') apply simp apply (auto simp add: mask_eq_exp_minus_1 min_def power_add algebra_simps dest!: le_Suc_ex) apply (metis add_minus_cancel le_add2 one_le_numeral power_add power_increasing uminus_add_conv_diff zle_diff1_eq) done lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (simp add: and_mask_bintr) apply (simp add: word_ubin.inverse_norm) apply (simp add: eq_mod_iff take_bit_eq_mod min_def) apply (fast intro!: lt2p_lem) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" - apply (simp add: dvd_eq_mod_eq_0 and_mask_mod_2p) - apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs del: word_of_int_0) - apply (subst word_uint.norm_Rep [symmetric]) - apply (simp only: bintrunc_bintrunc_min take_bit_eq_mod [symmetric] min_def) - apply auto - done + 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" - apply (simp flip: and_mask_dvd) - apply transfer - using dvd_nat_abs_iff [of _ \take_bit LENGTH('a) k\ for k] - apply simp - done + 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" - apply (unfold word_size word_less_alt word_numeral_alt) - apply (auto simp add: word_of_int_power_hom word_uint.eq_norm - simp del: word_of_int_numeral) - done + by transfer simp lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" for x :: "'a::len word" apply (simp add: and_mask_bintr) apply transfer apply (simp add: ac_simps) apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge mod_pos_pos_trivial mult.commute mult.left_neutral mult_zero_left not_le of_bool_def take_bit_eq_mod take_bit_nonnegative) done 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] - by (auto simp: and_mask_wi' word_of_int_homs word.abs_eq_iff take_bit_eq_mod mod_simps) + apply (auto simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + done 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] - by (auto simp: and_mask_wi' word_of_int_power_hom word.abs_eq_iff take_bit_eq_mod mod_simps) + apply (auto simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + done lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by transfer (simp add: take_bit_minus_one_eq_mask) 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 (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 exp_eq_zero_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 (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 slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" by (simp add: slice_shiftr nth_ucast nth_shiftr) lemma ucast_slice1: "ucast w = slice1 (size w) w" apply (simp add: slice1_def) apply transfer apply simp done 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\ apply (simp add: bit_slice1_iff bit_word_reverse_iff) using * ** apply (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) apply auto done 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)" apply (unfold slice_def word_size) apply (rule rev_slice1) apply arith done subsubsection \Revcast\ definition revcast :: \'a::len word \ 'b::len word\ where \revcast = slice1 LENGTH('b)\ lemma bit_revcast_iff: \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" apply auto apply (rule bit_word_eqI) apply (cases \LENGTH('a) \ LENGTH('b)\) apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le bit_imp_le_length) using bit_imp_le_length apply fastforce using bit_imp_le_length apply fastforce done 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 lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) done lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) done lemma cast_down_rev [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = (ucast w :: 'a::len word) << n" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) apply auto apply (metis add.commute add_diff_cancel_right) apply (metis diff_add_inverse2 diff_diff_add) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: \sint x + sint y = sint (x + y) \ (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ for x y :: \'a::len word\ proof - obtain n where n: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] by (auto intro: ccontr) have \sint x + sint y = sint (x + y) \ (sint (x + y) < 0 \ sint x < 0) \ (sint (x + y) < 0 \ sint y < 0)\ using sint_range' [of x] sint_range' [of y] apply (auto simp add: not_less) apply (unfold sint_word_ariths word_sbin.set_iff_norm [symmetric] sints_num) apply (auto simp add: signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) done then show ?thesis apply (simp only: One_nat_def word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) apply (simp add: bit_last_iff) done qed lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_eq lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" - by (simp add: word_rsplit_def word_ubin.eq_norm) + by (simp add: word_rsplit_def of_nat_take_bit) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] lemma test_bit_cat [OF refl]: "wc = word_cat a b \ wc !! n = (n < size wc \ (if n < size b then b !! n else a !! (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" apply (frule word_ubin.norm_Rep [THEN ssubst]) apply (drule bin_split_trunc1) apply (drule sym [THEN trans]) apply assumption apply safe done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" apply (unfold word_split_bin' test_bit_bin) apply (clarify) apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) - apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq ac_simps bin_nth_uint_imp) + apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_unsigned_iff ac_simps exp_eq_zero_iff dest: bit_imp_le_length) done lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. b !! n = (n < size b \ c !! n)) \ (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done \ \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 \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int (bin_cat w (size b) (uint b))" apply transfer using bintr_cat by auto lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" apply (rule word_eqI) apply (drule test_bit_split) apply (clarsimp simp add : test_bit_cat word_size) apply safe apply arith done lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" apply (drule test_bit_split) apply (rule conjI) apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ done lemma slice_cat1 [OF refl]: "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) done 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 a + size b >= size c \ word_cat a b = c" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) apply safe apply arith done lemma word_split_cat_alt: "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" apply (case_tac "word_split w") apply (rule trans, assumption) apply (drule test_bit_split) apply safe apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ done text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" by (simp add: word_rsplit_def bin_rsplit_all) lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) apply (rule_tac f = "\x. bin_nth x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) apply (simp add : word_ubin.eq_norm) apply (erule bin_rsplit_size_sign [OF len_gt_0 refl]) done 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: 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 lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def bin_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) (simp add: test_bit_eq_bit) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] lemma test_bit_rsplit_alt: \(word_rsplit w :: 'b::len word list) ! i !! m \ w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) apply (rule test_bit_cong) apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) apply simp apply (rule that(1)) apply simp apply (rule test_bit_rsplit) apply (rule refl) apply (rule asm_rl) apply (rule that(2)) apply (rule diff_Suc_less) apply (rule that(3)) done lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtrans(4) [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtrans(7), rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst rev_nth) apply arith apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (cases "size ws") apply simp_all done 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_rl [simp]: \word_rotl k (word_rotr k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done lemma word_rot_lr [simp]: \word_rotr k (word_rotl k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done 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 (simp add: minus_equation_iff [of \int m\]) apply (drule sym [of _ \int m\]) apply simp apply (metis add.commute add_minus_cancel diff_minus_eq_add len_gt_0 less_imp_of_nat_less less_nat_zero_code mod_mult_self3 of_nat_gt_0 zmod_minus1) apply (metis (no_types, hide_lams) 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] subsubsection \"Word rotation commutes with bit-wise operations\ \ \using locale to not pollute lemma namespace\ locale word_rotate 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" apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) done 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\ 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 (cases x rule: word_uint.Abs_cases) (simp add: uints_num) 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 (cases x rule: word_unat.Abs_cases) (simp add: unats_def) lemma max_word_max [intro!]: "n \ max_word" by (fact word_order.extremum) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" by (subst word_uint.Abs_norm [symmetric]) 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 = max_word" by (simp add: eq_neg_iff_add_eq_0) lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma word_and_max: "x AND max_word = x" by (fact word_log_esimps) lemma word_or_max: "x OR max_word = max_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 (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_or_not [simp]: "x OR NOT x = max_word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma shiftr_x_0 [iff]: "x >> 0 = x" for x :: "'a::len word" by transfer simp lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" by transfer simp lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) 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 map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto 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: word_arith_wis int_word_uint mod_add_if_z word_size) + apply (simp only: word_arith_wis word_size uint_word_of_int_eq) + apply (auto simp add: not_less take_bit_int_eq_self_iff) + apply (metis not_less uint_plus_if' word_add_def word_ubin.eq_norm) + done 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" apply (subst word_arith_nat_defs) apply (subst unat_of_nat) apply (auto simp add: not_less word_size) apply (metis not_le unat_plus_if' unat_word_ariths(1)) done lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" - by (simp add: word_gt_0) + 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: word_arith_wis int_word_uint mod_sub_if_z word_size) + apply (simp only: word_arith_wis word_size uint_word_of_int_eq) + apply (auto simp add: take_bit_int_eq_self_iff not_le) + apply (metis not_le uint_sub_if' word_sub_wi word_ubin.eq_norm) + done lemma unat_sub: \unat (a - b) = unat a - unat b\ if \b \ a\ proof - from that have \unat b \ unat a\ by transfer simp with that show ?thesis apply transfer apply simp apply (subst take_bit_diff [symmetric]) apply (subst nat_take_bit_eq) apply (simp add: nat_le_eq_zle) apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less bintr_lt2p) done qed 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)" proof - have *: "2^LENGTH('a) - i = -i + 2^LENGTH('a)" by simp show ?thesis apply (subst *) apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) apply simp done qed lemmas word_of_int_inj = word_uint.Abs_inject [unfolded uints_num, simplified] 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: "P 0 \ (\n. n < m \ P n \ P (1 + n)) \ P m" - for P :: "'a::len word \ bool" - apply (cases m) - apply atomize - apply (erule rev_mp)+ - apply (rule_tac x=m in spec) - apply (induct_tac n) - apply simp - apply clarsimp - apply (erule impE) - apply clarsimp - apply (erule_tac x=n in allE) - apply (erule impE) - apply (simp add: unat_arith_simps) - apply (clarsimp simp: unat_of_nat) - apply simp - apply (erule_tac x="of_nat na" in allE) - apply (erule impE) - apply (simp add: unat_arith_simps) - apply (clarsimp simp: unat_of_nat) - apply simp - done +lemma word_induct_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)\ + apply (rule Suc.IH) + apply (rule Suc.prems) + apply (erule less_trans) + apply (rule *) + apply assumption + done + 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 (erule word_induct_less) simp + by (rule word_induct_less) lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" - apply (rule word_induct) - apply simp - apply (case_tac "1 + n = 0") + apply (rule word_induct_less) + apply simp_all + apply (case_tac "1 + na = 0") apply auto done 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: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" apply (auto simp add: word_rec_def unat_word_ariths) apply (metis (mono_tags, lifting) old.nat.simps(7) unatSuc word_unat.Rep_inverse word_unat.eq_norm word_unat.td_th) done lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" apply (rule subst[where t="n" and s="1 + (n - 1)"]) apply simp apply (subst word_rec_Suc) apply simp apply simp done lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" apply (erule rev_mp) apply (rule_tac x=z in spec) apply (rule_tac x=f in spec) apply (induct n) apply (simp add: word_rec_0) apply clarsimp apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) apply simp apply (case_tac "1 + (n - m) = 0") apply (simp add: word_rec_0) apply (rule_tac f = "word_rec a b" for a b in arg_cong) apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) apply simp apply (simp (no_asm_use)) apply (simp add: word_rec_Suc word_rec_in2) apply (erule impE) apply uint_arith apply (drule_tac x="x \ (+) 1" in spec) apply (drule_tac x="x 0 xa" in spec) apply simp apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) apply (clarsimp simp add: fun_eq_iff) apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) apply simp apply (rule refl) apply (rule refl) done lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) (auto simp add: word_rec_0 word_rec_Suc) lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" apply (erule rev_mp) apply (induct n) apply (auto simp add: word_rec_0 word_rec_Suc) apply (drule spec, erule mp) apply uint_arith apply (drule_tac x=n in spec, erule impE) apply uint_arith apply simp done lemma word_rec_max: "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) apply simp apply simp apply (rule word_rec_id_eq) apply clarsimp apply (drule spec, rule mp, erule mp) apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) prefer 2 apply assumption apply simp apply (erule contrapos_pn) apply simp apply (drule arg_cong[where f="\x. x - n"]) apply simp done subsection \More\ lemma test_bit_1' [simp]: "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma mask_1: "mask 1 = 1" by transfer (simp add: min_def mask_Suc_exp) lemma mask_Suc_0: "mask (Suc 0) = 1" using mask_1 by simp lemma mask_numeral: "mask (numeral n) = 2 * mask (pred_numeral n) + (1 :: 'a::len word)" by (simp add: mask_Suc_rec numeral_eq_Suc) lemma bin_last_bintrunc: "bin_last (take_bit l n) = (l > 0 \ bin_last n)" by simp lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) lemma bintrunc_shiftl: "take_bit n (m << i) = take_bit (n - i) m << i" for m :: int by (rule bit_eqI) (auto simp add: bit_take_bit_iff) lemma uint_shiftl: "uint (n << i) = take_bit (size n) (uint n << i)" by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) subsection \Misc\ ML_file \Tools/word_lib.ML\ ML_file \Tools/smt_word.ML\ end diff --git a/src/HOL/ex/Bit_Lists.thy b/src/HOL/ex/Bit_Lists.thy --- a/src/HOL/ex/Bit_Lists.thy +++ b/src/HOL/ex/Bit_Lists.thy @@ -1,357 +1,357 @@ (* Author: Florian Haftmann, TUM *) section \Proof(s) of concept for algebraically founded lists of bits\ theory Bit_Lists imports - "HOL-Word.Conversions" "HOL-Library.More_List" + "HOL-Word.Word" "HOL-Library.More_List" begin subsection \Fragments of algebraic bit representations\ context comm_semiring_1 begin abbreviation (input) unsigned_of_bits :: "bool list \ 'a" where "unsigned_of_bits \ horner_sum of_bool 2" lemma unsigned_of_bits_replicate_False [simp]: "unsigned_of_bits (replicate n False) = 0" by (induction n) simp_all end context unique_euclidean_semiring_with_bit_shifts begin lemma unsigned_of_bits_append [simp]: "unsigned_of_bits (bs @ cs) = unsigned_of_bits bs + push_bit (length bs) (unsigned_of_bits cs)" by (induction bs) (simp_all add: push_bit_double, simp_all add: algebra_simps) lemma unsigned_of_bits_take [simp]: "unsigned_of_bits (take n bs) = take_bit n (unsigned_of_bits bs)" proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) (simp_all add: ac_simps take_bit_Suc) qed lemma unsigned_of_bits_drop [simp]: "unsigned_of_bits (drop n bs) = drop_bit n (unsigned_of_bits bs)" proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) (simp_all add: drop_bit_Suc) qed lemma bit_unsigned_of_bits_iff: \bit (unsigned_of_bits bs) n \ nth_default False bs n\ proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) (simp_all add: bit_Suc) qed primrec n_bits_of :: "nat \ 'a \ bool list" where "n_bits_of 0 a = []" | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)" lemma n_bits_of_eq_iff: "n_bits_of n a = n_bits_of n b \ take_bit n a = take_bit n b" apply (induction n arbitrary: a b) apply (auto elim!: evenE oddE simp add: take_bit_Suc mod_2_eq_odd) apply (metis dvd_triv_right even_plus_one_iff odd_iff_mod_2_eq_one) apply (metis dvd_triv_right even_plus_one_iff odd_iff_mod_2_eq_one) done lemma take_n_bits_of [simp]: "take m (n_bits_of n a) = n_bits_of (min m n) a" proof - define q and v and w where "q = min m n" and "v = m - q" and "w = n - q" then have "v = 0 \ w = 0" by auto then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a" by (induction q arbitrary: a) auto with q_def v_def w_def show ?thesis by simp qed lemma unsigned_of_bits_n_bits_of [simp]: "unsigned_of_bits (n_bits_of n a) = take_bit n a" by (induction n arbitrary: a) (simp_all add: ac_simps take_bit_Suc mod_2_eq_odd) end subsection \Syntactic bit representation\ class bit_representation = fixes bits_of :: "'a \ bool list" and of_bits :: "bool list \ 'a" assumes of_bits_of [simp]: "of_bits (bits_of a) = a" text \Unclear whether a \<^typ>\bool\ instantiation is needed or not\ instantiation nat :: bit_representation begin fun bits_of_nat :: "nat \ bool list" where "bits_of (n::nat) = (if n = 0 then [] else odd n # bits_of (n div 2))" lemma bits_of_nat_simps [simp]: "bits_of (0::nat) = []" "n > 0 \ bits_of n = odd n # bits_of (n div 2)" for n :: nat by simp_all declare bits_of_nat.simps [simp del] definition of_bits_nat :: "bool list \ nat" where [simp]: "of_bits_nat = unsigned_of_bits" \ \remove simp\ instance proof show "of_bits (bits_of n) = n" for n :: nat by (induction n rule: nat_bit_induct) simp_all qed end lemma bit_of_bits_nat_iff: \bit (of_bits bs :: nat) n \ nth_default False bs n\ by (simp add: bit_unsigned_of_bits_iff) lemma bits_of_Suc_0 [simp]: "bits_of (Suc 0) = [True]" by simp lemma bits_of_1_nat [simp]: "bits_of (1 :: nat) = [True]" by simp lemma bits_of_nat_numeral_simps [simp]: "bits_of (numeral Num.One :: nat) = [True]" (is ?One) "bits_of (numeral (Num.Bit0 n) :: nat) = False # bits_of (numeral n :: nat)" (is ?Bit0) "bits_of (numeral (Num.Bit1 n) :: nat) = True # bits_of (numeral n :: nat)" (is ?Bit1) proof - show ?One by simp define m :: nat where "m = numeral n" then have "m > 0" and *: "numeral n = m" "numeral (Num.Bit0 n) = 2 * m" "numeral (Num.Bit1 n) = Suc (2 * m)" by simp_all from \m > 0\ show ?Bit0 ?Bit1 by (simp_all add: *) qed lemma unsigned_of_bits_of_nat [simp]: "unsigned_of_bits (bits_of n) = n" for n :: nat using of_bits_of [of n] by simp instantiation int :: bit_representation begin fun bits_of_int :: "int \ bool list" where "bits_of_int k = odd k # (if k = 0 \ k = - 1 then [] else bits_of_int (k div 2))" lemma bits_of_int_simps [simp]: "bits_of (0 :: int) = [False]" "bits_of (- 1 :: int) = [True]" "k \ 0 \ k \ - 1 \ bits_of k = odd k # bits_of (k div 2)" for k :: int by simp_all lemma bits_of_not_Nil [simp]: "bits_of k \ []" for k :: int by simp declare bits_of_int.simps [simp del] definition of_bits_int :: "bool list \ int" where "of_bits_int bs = (if bs = [] \ \ last bs then unsigned_of_bits bs else unsigned_of_bits bs - 2 ^ length bs)" lemma of_bits_int_simps [simp]: "of_bits [] = (0 :: int)" "of_bits [False] = (0 :: int)" "of_bits [True] = (- 1 :: int)" "of_bits (bs @ [b]) = (unsigned_of_bits bs :: int) - (2 ^ length bs) * of_bool b" "of_bits (False # bs) = 2 * (of_bits bs :: int)" "bs \ [] \ of_bits (True # bs) = 1 + 2 * (of_bits bs :: int)" by (simp_all add: of_bits_int_def push_bit_of_1) instance proof show "of_bits (bits_of k) = k" for k :: int by (induction k rule: int_bit_induct) simp_all qed lemma bits_of_1_int [simp]: "bits_of (1 :: int) = [True, False]" by simp lemma bits_of_int_numeral_simps [simp]: "bits_of (numeral Num.One :: int) = [True, False]" (is ?One) "bits_of (numeral (Num.Bit0 n) :: int) = False # bits_of (numeral n :: int)" (is ?Bit0) "bits_of (numeral (Num.Bit1 n) :: int) = True # bits_of (numeral n :: int)" (is ?Bit1) "bits_of (- numeral (Num.Bit0 n) :: int) = False # bits_of (- numeral n :: int)" (is ?nBit0) "bits_of (- numeral (Num.Bit1 n) :: int) = True # bits_of (- numeral (Num.inc n) :: int)" (is ?nBit1) proof - show ?One by simp define k :: int where "k = numeral n" then have "k > 0" and *: "numeral n = k" "numeral (Num.Bit0 n) = 2 * k" "numeral (Num.Bit1 n) = 2 * k + 1" "numeral (Num.inc n) = k + 1" by (simp_all add: add_One) have "- (2 * k) div 2 = - k" "(- (2 * k) - 1) div 2 = - k - 1" by simp_all with \k > 0\ show ?Bit0 ?Bit1 ?nBit0 ?nBit1 by (simp_all add: *) qed lemma bit_of_bits_int_iff: \bit (of_bits bs :: int) n \ nth_default (bs \ [] \ last bs) bs n\ proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n; cases b; cases bs) (simp_all add: bit_Suc) qed lemma of_bits_append [simp]: "of_bits (bs @ cs) = of_bits bs + push_bit (length bs) (of_bits cs :: int)" if "bs \ []" "\ last bs" using that proof (induction bs rule: list_nonempty_induct) case (single b) then show ?case by simp next case (cons b bs) then show ?case by (cases b) (simp_all add: push_bit_double) qed lemma of_bits_replicate_False [simp]: "of_bits (replicate n False) = (0 :: int)" by (auto simp add: of_bits_int_def) lemma of_bits_drop [simp]: "of_bits (drop n bs) = drop_bit n (of_bits bs :: int)" if "n < length bs" using that proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc n) with Cons.prems have "bs \ []" by auto with Suc Cons.IH [of n] Cons.prems show ?thesis by (cases b) (simp_all add: drop_bit_Suc) qed qed end lemma unsigned_of_bits_eq_of_bits: "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)" by (simp add: of_bits_int_def) unbundle word.lifting instantiation word :: (len) bit_representation begin lift_definition bits_of_word :: "'a word \ bool list" is "n_bits_of LENGTH('a)" by (simp add: n_bits_of_eq_iff) lift_definition of_bits_word :: "bool list \ 'a word" is unsigned_of_bits . instance proof fix a :: "'a word" show "of_bits (bits_of a) = a" by transfer simp qed end lifting_update word.lifting lifting_forget word.lifting subsection \Bit representations with bit operations\ class semiring_bit_representation = semiring_bit_operations + bit_representation + assumes and_eq: "length bs = length cs \ of_bits bs AND of_bits cs = of_bits (map2 (\) bs cs)" and or_eq: "length bs = length cs \ of_bits bs OR of_bits cs = of_bits (map2 (\) bs cs)" and xor_eq: "length bs = length cs \ of_bits bs XOR of_bits cs = of_bits (map2 (\) bs cs)" and push_bit_eq: "push_bit n a = of_bits (replicate n False @ bits_of a)" and drop_bit_eq: "n < length (bits_of a) \ drop_bit n a = of_bits (drop n (bits_of a))" class ring_bit_representation = ring_bit_operations + semiring_bit_representation + assumes not_eq: "not = of_bits \ map Not \ bits_of" instance nat :: semiring_bit_representation by standard (simp_all add: bit_eq_iff bit_unsigned_of_bits_iff nth_default_map2 [of _ _ _ False False] bit_and_iff bit_or_iff bit_xor_iff) instance int :: ring_bit_representation proof { fix bs cs :: \bool list\ assume \length bs = length cs\ then have \cs = [] \ bs = []\ by auto with \length bs = length cs\ have \zip bs cs \ [] \ last (map2 (\) bs cs) \ (bs \ [] \ last bs) \ (cs \ [] \ last cs)\ and \zip bs cs \ [] \ last (map2 (\) bs cs) \ (bs \ [] \ last bs) \ (cs \ [] \ last cs)\ and \zip bs cs \ [] \ last (map2 (\) bs cs) \ ((bs \ [] \ last bs) \ (cs \ [] \ last cs))\ by (auto simp add: last_map last_zip zip_eq_Nil_iff prod_eq_iff) then show \of_bits bs AND of_bits cs = (of_bits (map2 (\) bs cs) :: int)\ and \of_bits bs OR of_bits cs = (of_bits (map2 (\) bs cs) :: int)\ and \of_bits bs XOR of_bits cs = (of_bits (map2 (\) bs cs) :: int)\ by (simp_all add: fun_eq_iff bit_eq_iff bit_and_iff bit_or_iff bit_xor_iff bit_not_iff bit_of_bits_int_iff \length bs = length cs\ nth_default_map2 [of bs cs _ \bs \ [] \ last bs\ \cs \ [] \ last cs\]) } show \push_bit n k = of_bits (replicate n False @ bits_of k)\ for k :: int and n :: nat by (cases "n = 0") simp_all show \drop_bit n k = of_bits (drop n (bits_of k))\ if \n < length (bits_of k)\ for k :: int and n :: nat using that by simp show \(not :: int \ _) = of_bits \ map Not \ bits_of\ proof (rule sym, rule ext) fix k :: int show \(of_bits \ map Not \ bits_of) k = NOT k\ by (induction k rule: int_bit_induct) (simp_all add: not_int_def) qed qed end