diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,1359 +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) + by (auto 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_trivial_iff: fixes i k :: int shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" proof - have "i mod k = i \ i div k = 0" 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_trivial_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 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_trivial_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_trivial_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 end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,2086 +1,2086 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (insert assms, simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto elim!: dvdE simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") (auto elim: dvdE) lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ lemma power_mod [mod_simps]: "((a mod b) ^ n) mod b = (a ^ n) mod b" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" by (simp add: mod_mult_right_eq) with Suc show ?case by (simp add: mod_mult_left_eq mod_mult_right_eq) qed lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp qed end class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel begin subclass idom_divide .. lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" using div_mult_mult1 [of "- 1" a b] by simp lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" using mod_mult_mult1 [of "- 1" a b] by simp lemma div_minus_right: "a div (- b) = (- a) div b" using div_minus_minus [of "- a" b] by simp lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" using mod_minus_minus [of "- a" b] by simp lemma div_minus1_right [simp]: "a div (- 1) = - a" using div_minus_right [of a 1] by simp lemma mod_minus1_right [simp]: "a mod (- 1) = 0" using mod_minus_right [of a 1] by simp text \Negation respects modular equivalence.\ lemma mod_minus_eq [mod_simps]: "(- (a mod b)) mod b = (- a) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: div_mult_mod_eq) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. then show ?thesis by (simp add: mod_minus_eq) qed text \Subtraction respects modular equivalence.\ lemma mod_diff_left_eq [mod_simps]: "(a mod c - b) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp lemma mod_diff_right_eq [mod_simps]: "(a - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_eq: "(a mod c - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp lemma minus_mod_self2 [simp]: "(a - b) mod b = a mod b" using mod_diff_right_eq [of a b b] by (simp add: mod_diff_right_eq) lemma minus_mod_self1 [simp]: "(b - a) mod b = - a mod b" using mod_add_self2 [of "- a" b] by simp lemma mod_eq_dvd_iff: "a mod c = b mod c \ c dvd a - b" (is "?P \ ?Q") proof assume ?P then have "(a mod c - b mod c) mod c = 0" by simp then show ?Q by (simp add: dvd_eq_mod_eq_0 mod_simps) next assume ?Q then obtain d where d: "a - b = c * d" .. then have "a = c * d + b" by (simp add: algebra_simps) then show ?P by simp qed lemma mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: "euclidean_size (a * b) = euclidean_size a * euclidean_size b" fixes division_segment :: "'a \ 'a" assumes is_unit_division_segment [simp]: "is_unit (division_segment a)" and division_segment_mult: "a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b" and division_segment_mod: "b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b" assumes div_bounded: "b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q" begin lemma division_segment_not_0 [simp]: "division_segment a \ 0" using is_unit_division_segment [of a] is_unitE [of "division_segment a"] by blast lemma divmod_cases [case_names divides remainder by0]: obtains (divides) q where "b \ 0" and "a div b = q" and "a mod b = 0" and "a = q * b" | (remainder) q r where "b \ 0" and "division_segment r = division_segment b" and "euclidean_size r < euclidean_size b" and "r \ 0" and "a div b = q" and "a mod b = r" and "a = q * b + r" | (by0) "b = 0" proof (cases "b = 0") case True then show thesis by (rule by0) next case False show thesis proof (cases "b dvd a") case True then obtain q where "a = b * q" .. with \b \ 0\ divides show thesis by (simp add: ac_simps) next case False then have "a mod b \ 0" by (simp add: mod_eq_0_iff_dvd) moreover from \b \ 0\ \\ b dvd a\ have "division_segment (a mod b) = division_segment b" by (rule division_segment_mod) moreover have "euclidean_size (a mod b) < euclidean_size b" using \b \ 0\ by (rule mod_size_less) moreover have "a = a div b * b + a mod b" by (simp add: div_mult_mod_eq) ultimately show thesis using \b \ 0\ by (blast intro!: remainder) qed qed lemma div_eqI: "a div b = q" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "(q * b + r) div b = q" by (auto intro: div_bounded) with that show ?thesis by simp qed lemma mod_eqI: "a mod b = r" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "a div b = q" by (rule div_eqI) moreover have "a div b * b + a mod b = a" by (fact div_mult_mod_eq) ultimately have "a div b * b + a mod b = a div b * b + r" using \q * b + r = a\ by simp then show ?thesis by simp qed subclass euclidean_semiring_cancel proof show "(a + c * b) div b = c + a div b" if "b \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 with \b \ 0\ show ?thesis by simp next case (divides q) then show ?thesis by (simp add: ac_simps) next case (remainder q r) then show ?thesis by (auto intro: div_eqI simp add: algebra_simps) qed next show"(c * a) div (c * b) = a div b" if "c \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 then show ?thesis by simp next case (divides q) with \c \ 0\ show ?thesis by (simp add: mult.left_commute [of c]) next case (remainder q r) from \b \ 0\ \c \ 0\ have "b * c \ 0" by simp from remainder \c \ 0\ have "division_segment (r * c) = division_segment (b * c)" and "euclidean_size (r * c) < euclidean_size (b * c)" by (simp_all add: division_segment_mult division_segment_mod euclidean_size_mult) with remainder show ?thesis by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps) (use \b * c \ 0\ in simp) qed qed lemma div_mult1_eq: "(a * b) div c = a * (b div c) + a * (b mod c) div c" proof (cases "a * (b mod c)" c rule: divmod_cases) case (divides q) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = (a * (b div c) + q) * c" using divides by (simp add: algebra_simps) finally have "(a * b) div c = \ div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = a * c * (b div c) + q * c + r" using remainder by (simp add: algebra_simps) finally show "(a * (b div c) + a * (b mod c) div c) * c + r = a * b" using remainder(5-7) by (simp add: algebra_simps) qed next case by0 then show ?thesis by simp qed lemma div_add1_eq: "(a + b) div c = a div c + b div c + (a mod c + b mod c) div c" proof (cases "a mod c + b mod c" c rule: divmod_cases) case (divides q) have "a + b = (a div c * c + a mod c) + (b div c * c + b mod c)" using mod_mult_div_eq [of a c] mod_mult_div_eq [of b c] by (simp add: ac_simps) also have "\ = (a div c + b div c) * c + (a mod c + b mod c)" by (simp add: algebra_simps) also have "\ = (a div c + b div c + q) * c" using divides by (simp add: algebra_simps) finally have "(a + b) div c = (a div c + b div c + q) * c div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "(a div c + b div c + q) * c + r + (a mod c + b mod c) = (a div c * c + a mod c) + (b div c * c + b mod c) + q * c + r" by (simp add: algebra_simps) also have "\ = a + b + (a mod c + b mod c)" by (simp add: div_mult_mod_eq remainder) (simp add: ac_simps) finally show "(a div c + b div c + (a mod c + b mod c) div c) * c + r = a + b" using remainder by simp qed next case by0 then show ?thesis by simp qed lemma div_eq_0_iff: "a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0" (is "_ \ ?P") if "division_segment a = division_segment b" proof assume ?P with that show "a div b = 0" by (cases "b = 0") (auto intro: div_eqI) next assume "a div b = 0" then have "a mod b = a" using div_mult_mod_eq [of a b] by simp with mod_size_less [of b a] show ?P by auto qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: "nat \ nat" where [simp]: "normalize = (id :: nat \ nat)" definition unit_factor_nat :: "nat \ nat" where "unit_factor n = (if n = 0 then 0 else 1 :: nat)" lemma unit_factor_simps [simp]: "unit_factor 0 = (0::nat)" "unit_factor (Suc n) = 1" by (simp_all add: unit_factor_nat_def) definition divide_nat :: "nat \ nat \ nat" where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \ m})" instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: "nat \ nat" where [simp]: "euclidean_size_nat = id" definition division_segment_nat :: "nat \ nat" where [simp]: "division_segment_nat n = 1" definition modulo_nat :: "nat \ nat \ nat" where "m mod n = m - (m div n * (n::nat))" instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_nat_eqI: "m div n = q" if "n * q \ m" and "m < n * Suc q" for m n q :: nat by (rule div_eqI [of _ "m - n * q"]) (use that in \simp_all add: algebra_simps\) lemma mod_nat_eqI: "m mod n = r" if "r < n" and "r \ m" and "n dvd m - r" for m n r :: nat by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \simp_all add: algebra_simps\) lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_eqI mod_eqI) lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all context fixes m n q :: nat begin private lemma eucl_rel_mult2: "m mod n + n * (m div n mod q) < n * q" if "n > 0" and "q > 0" proof - from \n > 0\ have "m mod n < n" by (rule mod_less_divisor) from \q > 0\ have "m div n mod q < q" by (rule mod_less_divisor) then obtain s where "q = Suc (m div n mod q + s)" by (blast dest: less_imp_Suc_add) moreover have "m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)" using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately show ?thesis by simp qed lemma div_mult2_eq: "m div (n * q) = (m div n) div q" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: div_eqI [of _ "n * (m div n mod q) + m mod n"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed lemma mod_mult2_eq: "m mod (n * q) = n * (m div n mod q) + m mod n" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: mod_eqI [of _ _ "(m div n) div q"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed end lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed text \A fact for the mutilated chess board\ lemma mod_Suc: "Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))" (is "_ = ?rhs") proof (cases "n = 0") case True then show ?thesis by simp next case False have "Suc m mod n = Suc (m mod n) mod n" by (simp add: mod_simps) also have "\ = ?rhs" using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq) finally show ?thesis . qed lemma Suc_times_mod_eq: "Suc (m * n) mod m = 1" if "Suc 0 < m" using that by (simp add: mod_Suc) lemma Suc_times_numeral_mod_eq [simp]: "Suc (numeral k * n) mod numeral k = 1" if "numeral k \ (1::nat)" by (rule Suc_times_mod_eq) (use that in simp) lemma Suc_div_le_mono [simp]: "m div n \ Suc m div n" by (simp add: div_le_mono) text \These lemmas collapse some needless occurrences of Suc: at least three Sucs, since two and fewer are rewritten back to Suc again! We already have some rules to simplify operands smaller than 3.\ lemma div_Suc_eq_div_add3 [simp]: "m div Suc (Suc (Suc n)) = m div (3 + n)" by (simp add: Suc3_eq_add_3) lemma mod_Suc_eq_mod_add3 [simp]: "m mod Suc (Suc (Suc n)) = m mod (3 + n)" by (simp add: Suc3_eq_add_3) lemma Suc_div_eq_add3_div: "Suc (Suc (Suc m)) div n = (3 + m) div n" by (simp add: Suc3_eq_add_3) lemma Suc_mod_eq_add3_mod: "Suc (Suc (Suc m)) mod n = (3 + m) mod n" by (simp add: Suc3_eq_add_3) lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v lemma (in field_char_0) of_nat_div: "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)" proof - have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)" unfolding of_nat_add by (cases "n = 0") simp_all then show ?thesis by simp qed text \An ``induction'' law for modulus arithmetic.\ lemma mod_induct [consumes 3, case_names step]: "P m" if "P n" and "n < p" and "m < p" and step: "\n. n < p \ P n \ P (Suc n mod p)" using \m < p\ proof (induct m) case 0 show ?case proof (rule ccontr) assume "\ P 0" from \n < p\ have "0 < p" by simp from \n < p\ obtain m where "0 < m" and "p = n + m" by (blast dest: less_imp_add_positive) with \P n\ have "P (p - m)" by simp moreover have "\ P (p - m)" using \0 < m\ proof (induct m) case 0 then show ?case by simp next case (Suc m) show ?case proof assume P: "P (p - Suc m)" with \\ P 0\ have "Suc m < p" by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma split_div: "P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))" (is "?P = ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P i" by simp with False show ?P by (auto intro: * [of "m mod n"]) qed qed lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma split_mod: "P (m mod n) \ (n = 0 \ P m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P j))" (is "?P \ ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P j" by simp with False show ?P by (auto intro: * [of _ "m div n"]) qed qed subsection \Euclidean division on \<^typ>\int\\ instantiation int :: normalization_semidom begin definition normalize_int :: "int \ int" where [simp]: "normalize = (abs :: int \ int)" definition unit_factor_int :: "int \ int" where [simp]: "unit_factor = (sgn :: int \ int)" definition divide_int :: "int \ int \ int" where "k div l = (if l = 0 then 0 else if sgn k = sgn l then int (nat \k\ div nat \l\) else - int (nat \k\ div nat \l\ + of_bool (\ l dvd k)))" lemma divide_int_unfold: "(sgn k * int m) div (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ n = 0 then 0 else if sgn k = sgn l then int (m div n) else - int (m div n + of_bool (\ n dvd m)))" by (auto simp add: divide_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult nat_mult_distrib) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation int :: idom_modulo begin definition modulo_int :: "int \ int \ int" where "k mod l = (if l = 0 then k else if sgn k = sgn l then sgn l * int (nat \k\ mod nat \l\) else sgn l * (\l\ * of_bool (\ l dvd k) - int (nat \k\ mod nat \l\)))" lemma modulo_int_unfold: "(sgn k * int m) mod (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ n = 0 then sgn k * int m else if sgn k = sgn l then sgn l * int (m mod n) else sgn l * (int (n * of_bool (\ n dvd m)) - int (m mod n)))" by (auto simp add: modulo_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult nat_mult_distrib) instance proof fix k l :: int obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (auto simp add: divide_int_unfold modulo_int_unfold algebra_simps dest!: sgn_not_eq_imp) (simp_all add: of_nat_mult [symmetric] of_nat_add [symmetric] distrib_left [symmetric] minus_mult_right del: of_nat_mult minus_mult_right [symmetric]) qed end instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg abs_mult mod_greater_zero_iff_not_dvd) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis - by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg - sgn_mult mod_eq_0_iff_dvd) + by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg sgn_mult) + (simp add: dvd_eq_mod_eq_0) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: q l divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) end lemma pos_mod_bound [simp]: "k mod l < l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_pos_neg_trivial: "k div l = - 1" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule div_eqI [of _ "k + l"]) using False that apply (simp_all add: division_segment_int_def) done qed lemma mod_pos_neg_trivial: "k mod l = k + l" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule mod_eqI [of _ _ \- 1\]) using False that apply (simp_all add: division_segment_int_def) done qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ text \Distributive laws for function \nat\.\ lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" proof (cases a "of_nat m * of_nat n" rule: divmod_cases) case (divides q) then show ?thesis using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] by (simp add: ac_simps) next case (remainder q r) then have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by simp with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with remainder(6) have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with remainder(1) have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp with remainder(5) remainder(7) show ?thesis using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] by (simp add: ac_simps) next case by0 then show ?thesis by auto qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD min.absorb2) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def) subsection \Code generation\ code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Groups_Big.thy b/src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy +++ b/src/HOL/Groups_Big.thy @@ -1,1654 +1,1673 @@ (* Title: HOL/Groups_Big.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Big sum and product over finite (non-empty) sets\ theory Groups_Big imports Power begin subsection \Generic monoid operation over a set\ locale comm_monoid_set = comm_monoid begin subsubsection \Standard sum or product indexed by a finite set\ interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) definition F :: "('b \ 'a) \ 'b set \ 'a" where eq_fold: "F g A = Finite_Set.fold (f \ g) \<^bold>1 A" lemma infinite [simp]: "\ finite A \ F g A = \<^bold>1" by (simp add: eq_fold) lemma empty [simp]: "F g {} = \<^bold>1" by (simp add: eq_fold) lemma insert [simp]: "finite A \ x \ A \ F g (insert x A) = g x \<^bold>* F g A" by (simp add: eq_fold) lemma remove: assumes "finite A" and "x \ A" shows "F g A = g x \<^bold>* F g (A - {x})" proof - from \x \ A\ obtain B where B: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ B have "finite B" by simp ultimately show ?thesis by simp qed lemma insert_remove: "finite A \ F g (insert x A) = g x \<^bold>* F g (A - {x})" by (cases "x \ A") (simp_all add: remove insert_absorb) lemma insert_if: "finite A \ F g (insert x A) = (if x \ A then F g A else g x \<^bold>* F g A)" by (cases "x \ A") (simp_all add: insert_absorb) lemma neutral: "\x\A. g x = \<^bold>1 \ F g A = \<^bold>1" by (induct A rule: infinite_finite_induct) simp_all lemma neutral_const [simp]: "F (\_. \<^bold>1) A = \<^bold>1" by (simp add: neutral) lemma union_inter: assumes "finite A" and "finite B" shows "F g (A \ B) \<^bold>* F g (A \ B) = F g A \<^bold>* F g B" \ \The reversed orientation looks more natural, but LOOPS as a simprule!\ using assms proof (induct A) case empty then show ?case by simp next case (insert x A) then show ?case by (auto simp: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) qed corollary union_inter_neutral: assumes "finite A" and "finite B" and "\x \ A \ B. g x = \<^bold>1" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter [symmetric] neutral) corollary union_disjoint: assumes "finite A" and "finite B" assumes "A \ B = {}" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter_neutral) lemma union_diff2: assumes "finite A" and "finite B" shows "F g (A \ B) = F g (A - B) \<^bold>* F g (B - A) \<^bold>* F g (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst union_disjoint, auto)+ qed lemma subset_diff: assumes "B \ A" and "finite A" shows "F g A = F g (A - B) \<^bold>* F g B" proof - from assms have "finite (A - B)" by auto moreover from assms have "finite B" by (rule finite_subset) moreover from assms have "(A - B) \ B = {}" by auto ultimately have "F g (A - B \ B) = F g (A - B) \<^bold>* F g B" by (rule union_disjoint) moreover from assms have "A \ B = A" by auto ultimately show ?thesis by simp qed lemma Int_Diff: assumes "finite A" shows "F g A = F g (A \ B) \<^bold>* F g (A - B)" by (subst subset_diff [where B = "A - B"]) (auto simp: Diff_Diff_Int assms) lemma setdiff_irrelevant: assumes "finite A" shows "F g (A - {x. g x = z}) = F g A" using assms by (induct A) (simp_all add: insert_Diff_if) lemma not_neutral_contains_not_neutral: assumes "F g A \ \<^bold>1" obtains a where "a \ A" and "g a \ \<^bold>1" proof - from assms have "\a\A. g a \ \<^bold>1" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then show ?case by fastforce qed with that show thesis by blast qed lemma reindex: assumes "inj_on h A" shows "F g (h ` A) = F (g \ h) A" proof (cases "finite A") case True with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) next case False with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) with False show ?thesis by simp qed lemma cong [fundef_cong]: assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "F g A = F h B" using g_h unfolding \A = B\ by (induct B rule: infinite_finite_induct) auto lemma cong_simp [cong]: "\ A = B; \x. x \ B =simp=> g x = h x \ \ F (\x. g x) A = F (\x. h x) B" by (rule cong) (simp_all add: simp_implies_def) lemma reindex_cong: assumes "inj_on l B" assumes "A = l ` B" assumes "\x. x \ B \ g (l x) = h x" shows "F g A = F h B" using assms by (simp add: reindex) lemma image_eq: assumes "inj_on g A" shows "F (\x. x) (g ` A) = F g A" using assms reindex_cong by fastforce lemma UNION_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "F g (\(A ` I)) = F (\x. F g (A x)) I" using assms proof (induction rule: finite_induct) case (insert i I) then have "\j\I. j \ i" by blast with insert.prems have "A i \ \(A ` I) = {}" by blast with insert show ?case by (simp add: union_disjoint) qed auto lemma Union_disjoint: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A \ B = {}" shows "F g (\C) = (F \ F) g C" proof (cases "finite C") case True from UNION_disjoint [OF this assms] show ?thesis by simp next case False then show ?thesis by (auto dest: finite_UnionD intro: infinite) qed lemma distrib: "F (\x. g x \<^bold>* h x) A = F g A \<^bold>* F h A" by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) lemma Sigma: assumes "finite A" "\x\A. finite (B x)" shows "F (\x. F (g x) (B x)) A = F (case_prod g) (SIGMA x:A. B x)" unfolding Sigma_def proof (subst UNION_disjoint) show "F (\x. F (g x) (B x)) A = F (\x. F (\(x, y). g x y) (\y\B x. {(x, y)})) A" proof (rule cong [OF refl]) show "F (g x) (B x) = F (\(x, y). g x y) (\y\B x. {(x, y)})" if "x \ A" for x using that assms by (simp add: UNION_disjoint) qed qed (use assms in auto) lemma related: assumes Re: "R \<^bold>1 \<^bold>1" and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 \<^bold>* y1) (x2 \<^bold>* y2)" and fin: "finite S" and R_h_g: "\x\S. R (h x) (g x)" shows "R (F h S) (F g S)" using fin by (rule finite_subset_induct) (use assms in auto) lemma mono_neutral_cong_left: assumes "finite T" and "S \ T" and "\i \ T - S. h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "F g S = F h T" proof- have eq: "T = S \ (T - S)" using \S \ T\ by blast have d: "S \ (T - S) = {}" using \S \ T\ by blast from \finite T\ \S \ T\ have f: "finite S" "finite (T - S)" by (auto intro: finite_subset) show ?thesis using assms(4) by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)]) qed lemma mono_neutral_cong_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ F g T = F h S" by (auto intro!: mono_neutral_cong_left [symmetric]) lemma mono_neutral_left: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g S = F g T" by (blast intro: mono_neutral_cong_left) lemma mono_neutral_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g T = F g S" by (blast intro!: mono_neutral_left [symmetric]) lemma mono_neutral_cong: assumes [simp]: "finite T" "finite S" and *: "\i. i \ T - S \ h i = \<^bold>1" "\i. i \ S - T \ g i = \<^bold>1" and gh: "\x. x \ S \ T \ g x = h x" shows "F g S = F h T" proof- have "F g S = F g (S \ T)" by(rule mono_neutral_right)(auto intro: *) also have "\ = F h (S \ T)" using refl gh by(rule cong) also have "\ = F h T" by(rule mono_neutral_left)(auto intro: *) finally show ?thesis . qed lemma reindex_bij_betw: "bij_betw h S T \ F (\x. g (h x)) S = F g T" by (auto simp: bij_betw_def reindex) lemma reindex_bij_witness: assumes witness: "\a. a \ S \ i (j a) = a" "\a. a \ S \ j a \ T" "\b. b \ T \ j (i b) = b" "\b. b \ T \ i b \ S" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have "bij_betw j S T" using bij_betw_byWitness[where A=S and f=j and f'=i and A'=T] witness by auto moreover have "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) ultimately show ?thesis by (simp add: reindex_bij_betw) qed lemma reindex_bij_betw_not_neutral: assumes fin: "finite S'" "finite T'" assumes bij: "bij_betw h (S - S') (T - T')" assumes nn: "\a. a \ S' \ g (h a) = z" "\b. b \ T' \ g b = z" shows "F (\x. g (h x)) S = F g T" proof - have [simp]: "finite S \ finite T" using bij_betw_finite[OF bij] fin by auto show ?thesis proof (cases "finite S") case True with nn have "F (\x. g (h x)) S = F (\x. g (h x)) (S - S')" by (intro mono_neutral_cong_right) auto also have "\ = F g (T - T')" using bij by (rule reindex_bij_betw) also have "\ = F g T" using nn \finite S\ by (intro mono_neutral_cong_left) auto finally show ?thesis . next case False then show ?thesis by simp qed qed lemma reindex_nontrivial: assumes "finite A" and nz: "\x y. x \ A \ y \ A \ x \ y \ h x = h y \ g (h x) = \<^bold>1" shows "F g (h ` A) = F (g \ h) A" proof (subst reindex_bij_betw_not_neutral [symmetric]) show "bij_betw h (A - {x \ A. (g \ h) x = \<^bold>1}) (h ` A - h ` {x \ A. (g \ h) x = \<^bold>1})" using nz by (auto intro!: inj_onI simp: bij_betw_def) qed (use \finite A\ in auto) lemma reindex_bij_witness_not_neutral: assumes fin: "finite S'" "finite T'" assumes witness: "\a. a \ S - S' \ i (j a) = a" "\a. a \ S - S' \ j a \ T - T'" "\b. b \ T - T' \ j (i b) = b" "\b. b \ T - T' \ i b \ S - S'" assumes nn: "\a. a \ S' \ g a = z" "\b. b \ T' \ h b = z" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have bij: "bij_betw j (S - (S' \ S)) (T - (T' \ T))" using witness by (intro bij_betw_byWitness[where f'=i]) auto have F_eq: "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) show ?thesis unfolding F_eq using fin nn eq by (intro reindex_bij_betw_not_neutral[OF _ _ bij]) auto qed lemma delta_remove: assumes fS: "finite S" shows "F (\k. if k = a then b k else c k) S = (if a \ S then b a \<^bold>* F c (S-{a}) else F c (S-{a}))" proof - let ?f = "(\k. if k = a then b k else c k)" show ?thesis proof (cases "a \ S") case False then have "\k\S. ?f k = c k" by simp with False show ?thesis by simp next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have "F ?f S = F ?f ?A \<^bold>* F ?f ?B" using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] by simp with True show ?thesis using comm_monoid_set.remove comm_monoid_set_axioms fS by fastforce qed qed lemma delta [simp]: assumes fS: "finite S" shows "F (\k. if k = a then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" by (simp add: delta_remove [OF assms]) lemma delta' [simp]: assumes fin: "finite S" shows "F (\k. if a = k then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" using delta [OF fin, of a b, symmetric] by (auto intro: cong) lemma If_cases: fixes P :: "'b \ bool" and g h :: "'b \ 'a" assumes fin: "finite A" shows "F (\x. if P x then h x else g x) A = F h (A \ {x. P x}) \<^bold>* F g (A \ - {x. P x})" proof - have a: "A = A \ {x. P x} \ A \ -{x. P x}" "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" by blast+ from fin have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto let ?g = "\x. if P x then h x else g x" from union_disjoint [OF f a(2), of ?g] a(1) show ?thesis by (subst (1 2) cong) simp_all qed lemma cartesian_product: "F (\x. F (g x) B) A = F (case_prod g) (A \ B)" proof (cases "A = {} \ B = {}") case True then show ?thesis by auto next case False then have "A \ {}" "B \ {}" by auto show ?thesis proof (cases "finite A \ finite B") case True then show ?thesis by (simp add: Sigma) next case False then consider "infinite A" | "infinite B" by auto then have "infinite (A \ B)" by cases (use \A \ {}\ \B \ {}\ in \auto dest: finite_cartesian_productD1 finite_cartesian_productD2\) then show ?thesis using False by auto qed qed lemma inter_restrict: assumes "finite A" shows "F g (A \ B) = F (\x. if x \ B then g x else \<^bold>1) A" proof - let ?g = "\x. if x \ A \ B then g x else \<^bold>1" have "\i\A - A \ B. (if i \ A \ B then g i else \<^bold>1) = \<^bold>1" by simp moreover have "A \ B \ A" by blast ultimately have "F ?g (A \ B) = F ?g A" using \finite A\ by (intro mono_neutral_left) auto then show ?thesis by simp qed lemma inter_filter: "finite A \ F g {x \ A. P x} = F (\x. if P x then g x else \<^bold>1) A" by (simp add: inter_restrict [symmetric, of A "{x. P x}" g, simplified mem_Collect_eq] Int_def) lemma Union_comp: assumes "\A \ B. finite A" and "\A1 A2 x. A1 \ B \ A2 \ B \ A1 \ A2 \ x \ A1 \ x \ A2 \ g x = \<^bold>1" shows "F g (\B) = (F \ F) g B" using assms proof (induct B rule: infinite_finite_induct) case (infinite A) then have "\ finite (\A)" by (blast dest: finite_UnionD) with infinite show ?case by simp next case empty then show ?case by simp next case (insert A B) then have "finite A" "finite B" "finite (\B)" "A \ B" and "\x\A \ \B. g x = \<^bold>1" and H: "F g (\B) = (F \ F) g B" by auto then have "F g (A \ \B) = F g A \<^bold>* F g (\B)" by (simp add: union_inter_neutral) with \finite B\ \A \ B\ show ?case by (simp add: H) qed lemma swap: "F (\i. F (g i) B) A = F (\j. F (\i. g i j) A) B" unfolding cartesian_product by (rule reindex_bij_witness [where i = "\(i, j). (j, i)" and j = "\(i, j). (j, i)"]) auto lemma swap_restrict: "finite A \ finite B \ F (\x. F (g x) {y. y \ B \ R x y}) A = F (\y. F (\x. g x y) {x. x \ A \ R x y}) B" by (simp add: inter_filter) (rule swap) lemma image_gen: assumes fin: "finite S" shows "F h S = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" proof - have "{y. y\ g`S \ g x = y} = {g x}" if "x \ S" for x using that by auto then have "F h S = F (\x. F (\y. h x) {y. y\ g`S \ g x = y}) S" by simp also have "\ = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" by (rule swap_restrict [OF fin finite_imageI [OF fin]]) finally show ?thesis . qed lemma group: assumes fS: "finite S" and fT: "finite T" and fST: "g ` S \ T" shows "F (\y. F h {x. x \ S \ g x = y}) T = F h S" unfolding image_gen[OF fS, of h g] by (auto intro: neutral mono_neutral_right[OF fT fST]) lemma Plus: fixes A :: "'b set" and B :: "'c set" assumes fin: "finite A" "finite B" shows "F g (A <+> B) = F (g \ Inl) A \<^bold>* F (g \ Inr) B" proof - have "A <+> B = Inl ` A \ Inr ` B" by auto moreover from fin have "finite (Inl ` A)" "finite (Inr ` B)" by auto moreover have "Inl ` A \ Inr ` B = {}" by auto moreover have "inj_on Inl A" "inj_on Inr B" by (auto intro: inj_onI) ultimately show ?thesis using fin by (simp add: union_disjoint reindex) qed lemma same_carrier: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" shows "F g A = F h B \ F g C = F h C" proof - have "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)" using \finite C\ subset by (auto elim: finite_subset) from subset have [simp]: "A - (C - A) = A" by auto from subset have [simp]: "B - (C - B) = B" by auto from subset have "C = A \ (C - A)" by auto then have "F g C = F g (A \ (C - A))" by simp also have "\ = F g (A - (C - A)) \<^bold>* F g (C - A - A) \<^bold>* F g (A \ (C - A))" using \finite A\ \finite (C - A)\ by (simp only: union_diff2) finally have *: "F g C = F g A" using trivial by simp from subset have "C = B \ (C - B)" by auto then have "F h C = F h (B \ (C - B))" by simp also have "\ = F h (B - (C - B)) \<^bold>* F h (C - B - B) \<^bold>* F h (B \ (C - B))" using \finite B\ \finite (C - B)\ by (simp only: union_diff2) finally have "F h C = F h B" using trivial by simp with * show ?thesis by simp qed lemma same_carrierI: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" assumes "F g C = F h C" shows "F g A = F h B" using assms same_carrier [of C A B] by simp lemma eq_general: assumes B: "\y. y \ B \ \!x. x \ A \ h x = y" and A: "\x. x \ A \ h x \ B \ \(h x) = \ x" shows "F \ A = F \ B" proof - have eq: "B = h ` A" by (auto dest: assms) have h: "inj_on h A" using assms by (blast intro: inj_onI) have "F \ A = F (\ \ h) A" using A by auto also have "\ = F \ B" by (simp add: eq reindex h) finally show ?thesis . qed lemma eq_general_inverses: assumes B: "\y. y \ B \ k y \ A \ h(k y) = y" and A: "\x. x \ A \ h x \ B \ k(h x) = x \ \(h x) = \ x" shows "F \ A = F \ B" by (rule eq_general [where h=h]) (force intro: dest: A B)+ subsubsection \HOL Light variant: sum/product indexed by the non-neutral subset\ text \NB only a subset of the properties above are proved\ definition G :: "['b \ 'a,'b set] \ 'a" where "G p I \ if finite {x \ I. p x \ \<^bold>1} then F p {x \ I. p x \ \<^bold>1} else \<^bold>1" lemma finite_Collect_op: shows "\finite {i \ I. x i \ \<^bold>1}; finite {i \ I. y i \ \<^bold>1}\ \ finite {i \ I. x i \<^bold>* y i \ \<^bold>1}" apply (rule finite_subset [where B = "{i \ I. x i \ \<^bold>1} \ {i \ I. y i \ \<^bold>1}"]) using left_neutral by force+ lemma empty' [simp]: "G p {} = \<^bold>1" by (auto simp: G_def) lemma eq_sum [simp]: "finite I \ G p I = F p I" by (auto simp: G_def intro: mono_neutral_cong_left) lemma insert' [simp]: assumes "finite {x \ I. p x \ \<^bold>1}" shows "G p (insert i I) = (if i \ I then G p I else p i \<^bold>* G p I)" proof - have "{x. x = i \ p x \ \<^bold>1 \ x \ I \ p x \ \<^bold>1} = (if p i = \<^bold>1 then {x \ I. p x \ \<^bold>1} else insert i {x \ I. p x \ \<^bold>1})" by auto then show ?thesis using assms by (simp add: G_def conj_disj_distribR insert_absorb) qed lemma distrib_triv': assumes "finite I" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" by (simp add: assms local.distrib) lemma non_neutral': "G g {x \ I. g x \ \<^bold>1} = G g I" by (simp add: G_def) lemma distrib': assumes "finite {x \ I. g x \ \<^bold>1}" "finite {x \ I. h x \ \<^bold>1}" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" proof - have "a \<^bold>* a \ a \ a \ \<^bold>1" for a by auto then have "G (\i. g i \<^bold>* h i) I = G (\i. g i \<^bold>* h i) ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1})" using assms by (force simp: G_def finite_Collect_op intro!: mono_neutral_cong) also have "\ = G g I \<^bold>* G h I" proof - have "F g ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G g I" "F h ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G h I" by (auto simp: G_def assms intro: mono_neutral_right) then show ?thesis using assms by (simp add: distrib) qed finally show ?thesis . qed lemma cong': assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "G g A = G h B" using assms by (auto simp: G_def cong: conj_cong intro: cong) lemma mono_neutral_cong_left': assumes "S \ T" and "\i. i \ T - S \ h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "G g S = G h T" proof - have *: "{x \ S. g x \ \<^bold>1} = {x \ T. h x \ \<^bold>1}" using assms by (metis DiffI subset_eq) then have "finite {x \ S. g x \ \<^bold>1} = finite {x \ T. h x \ \<^bold>1}" by simp then show ?thesis using assms by (auto simp add: G_def * intro: cong) qed lemma mono_neutral_cong_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ G g T = G h S" by (auto intro!: mono_neutral_cong_left' [symmetric]) lemma mono_neutral_left': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g S = G g T" by (blast intro: mono_neutral_cong_left') lemma mono_neutral_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g T = G g S" by (blast intro!: mono_neutral_left' [symmetric]) end subsection \Generalized summation over a set\ context comm_monoid_add begin sublocale sum: comm_monoid_set plus 0 defines sum = sum.F and sum' = sum.G .. abbreviation Sum ("\") where "\ \ sum (\x. x)" end text \Now: lots of fancy syntax. First, \<^term>\sum (\x. e) A\ is written \\x\A. e\.\ syntax (ASCII) "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(3SUM (_/:_)./ _)" [0, 51, 10] 10) syntax "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" \ "CONST sum (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10) syntax "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST sum (\x. t) {x. P}" print_translation \ let fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\Collect\, _) $ Abs (y, Ty, P)] = if x <> y then raise Match else let val x' = Syntax_Trans.mark_bound_body (x, Tx); val t' = subst_bound (x', t); val P' = subst_bound (x', P); in Syntax.const \<^syntax_const>\_qsum\ $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' end | sum_tr' _ = raise Match; in [(\<^const_syntax>\sum\, K sum_tr')] end \ subsubsection \Properties in more restricted classes of structures\ lemma sum_Un: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'b \ 'a::ab_group_add" by (subst sum.union_inter [symmetric]) (auto simp add: algebra_simps) lemma sum_Un2: assumes "finite (A \ B)" shows "sum f (A \ B) = sum f (A - B) + sum f (B - A) + sum f (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst sum.union_disjoint, auto)+ qed lemma sum_diff1: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" shows "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" using assms by induct (auto simp: insert_Diff_if) lemma sum_diff: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" "B \ A" shows "sum f (A - B) = sum f A - sum f B" proof - from assms(2,1) have "finite B" by (rule finite_subset) from this \B \ A\ show ?thesis proof induct case empty thus ?case by simp next case (insert x F) with \finite A\ \finite B\ show ?case by (simp add: Diff_insert[where a=x and B=F] sum_diff1 insert_absorb) qed qed lemma sum_diff1'_aux: fixes f :: "'a \ 'b::ab_group_add" assumes "finite F" "{i \ I. f i \ 0} \ F" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" using assms proof induct case (insert x F) have 1: "finite {x \ I. f x \ 0} \ finite {x \ I. x \ i \ f x \ 0}" by (erule rev_finite_subset) auto have 2: "finite {x \ I. x \ i \ f x \ 0} \ finite {x \ I. f x \ 0}" apply (drule finite_insert [THEN iffD2]) by (erule rev_finite_subset) auto have 3: "finite {i \ I. f i \ 0}" using finite_subset insert by blast show ?case using insert sum_diff1 [of "{i \ I. f i \ 0}" f i] by (auto simp: sum.G_def 1 2 3 set_diff_eq conj_ac) qed (simp add: sum.G_def) lemma sum_diff1': fixes f :: "'a \ 'b::ab_group_add" assumes "finite {i \ I. f i \ 0}" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" by (rule sum_diff1'_aux [OF assms order_refl]) lemma (in ordered_comm_monoid_add) sum_mono: "(\i. i\K \ f i \ g i) \ (\i\K. f i) \ (\i\K. g i)" by (induct K rule: infinite_finite_induct) (use add_mono in auto) lemma (in strict_ordered_comm_monoid_add) sum_strict_mono: assumes "finite A" "A \ {}" and "\x. x \ A \ f x < g x" shows "sum f A < sum g A" using assms proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case insert then show ?case by (auto simp: add_strict_mono) qed lemma sum_strict_mono_ex1: fixes f g :: "'i \ 'a::ordered_cancel_comm_monoid_add" assumes "finite A" and "\x\A. f x \ g x" and "\a\A. f a < g a" shows "sum f A < sum g A" proof- from assms(3) obtain a where a: "a \ A" "f a < g a" by blast have "sum f A = sum f ((A - {a}) \ {a})" by(simp add: insert_absorb[OF \a \ A\]) also have "\ = sum f (A - {a}) + sum f {a}" using \finite A\ by(subst sum.union_disjoint) auto also have "sum f (A - {a}) \ sum g (A - {a})" by (rule sum_mono) (simp add: assms(2)) also from a have "sum f {a} < sum g {a}" by simp also have "sum g (A - {a}) + sum g {a} = sum g((A - {a}) \ {a})" using \finite A\ by (subst sum.union_disjoint[symmetric]) auto also have "\ = sum g A" by (simp add: insert_absorb[OF \a \ A\]) finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono) qed lemma sum_mono_inv: fixes f g :: "'i \ 'a :: ordered_cancel_comm_monoid_add" assumes eq: "sum f I = sum g I" assumes le: "\i. i \ I \ f i \ g i" assumes i: "i \ I" assumes I: "finite I" shows "f i = g i" proof (rule ccontr) assume "\ ?thesis" with le[OF i] have "f i < g i" by simp with i have "\i\I. f i < g i" .. from sum_strict_mono_ex1[OF I _ this] le have "sum f I < sum g I" by blast with eq show False by simp qed lemma member_le_sum: fixes f :: "_ \ 'b::{semiring_1, ordered_comm_monoid_add}" assumes "i \ A" and le: "\x. x \ A - {i} \ 0 \ f x" and "finite A" shows "f i \ sum f A" proof - have "f i \ sum f (A \ {i})" by (simp add: assms) also have "... = (\x\A. if x \ {i} then f x else 0)" using assms sum.inter_restrict by blast also have "... \ sum f A" apply (rule sum_mono) apply (auto simp: le) done finally show ?thesis . qed lemma sum_negf: "(\x\A. - f x) = - (\x\A. f x)" for f :: "'b \ 'a::ab_group_add" by (induct A rule: infinite_finite_induct) auto lemma sum_subtractf: "(\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'b \'a::ab_group_add" using sum.distrib [of f "- g" A] by (simp add: sum_negf) lemma sum_subtractf_nat: "(\x. x \ A \ g x \ f x) \ (\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'a \ nat" by (induct A rule: infinite_finite_induct) (auto simp: sum_mono) context ordered_comm_monoid_add begin lemma sum_nonneg: "(\x. x \ A \ 0 \ f x) \ 0 \ sum f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "0 + 0 \ f x + sum f F" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonpos: "(\x. x \ A \ f x \ 0) \ sum f A \ 0" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "f x + sum f F \ 0 + 0" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonneg_eq_0_iff: "finite A \ (\x. x \ A \ 0 \ f x) \ sum f A = 0 \ (\x\A. f x = 0)" by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff sum_nonneg) lemma sum_nonneg_0: "finite s \ (\i. i \ s \ f i \ 0) \ (\ i \ s. f i) = 0 \ i \ s \ f i = 0" by (simp add: sum_nonneg_eq_0_iff) lemma sum_nonneg_leq_bound: assumes "finite s" "\i. i \ s \ f i \ 0" "(\i \ s. f i) = B" "i \ s" shows "f i \ B" proof - from assms have "f i \ f i + (\i \ s - {i}. f i)" by (intro add_increasing2 sum_nonneg) auto also have "\ = B" using sum.remove[of s i f] assms by simp finally show ?thesis by auto qed lemma sum_mono2: assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" shows "sum f A \ sum f B" proof - have "sum f A \ sum f A + sum f (B-A)" by (auto intro: add_increasing2 [OF sum_nonneg] nn) also from fin finite_subset[OF sub fin] have "\ = sum f (A \ (B-A))" by (simp add: sum.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma sum_le_included: assumes "finite s" "finite t" and "\y\t. 0 \ g y" "(\x\s. \y\t. i y = x \ f x \ g y)" shows "sum f s \ sum g t" proof - have "sum f s \ sum (\y. sum g {x. x\t \ i x = y}) s" proof (rule sum_mono) fix y assume "y \ s" with assms obtain z where z: "z \ t" "y = i z" "f y \ g z" by auto with assms show "f y \ sum g {x \ t. i x = y}" (is "?A y \ ?B y") using order_trans[of "?A (i z)" "sum g {z}" "?B (i z)", intro] by (auto intro!: sum_mono2) qed also have "\ \ sum (\y. sum g {x. x\t \ i x = y}) (i ` t)" using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg) also have "\ \ sum g t" using assms by (auto simp: sum.image_gen[symmetric]) finally show ?thesis . qed end lemma (in canonically_ordered_monoid_add) sum_eq_0_iff [simp]: "finite F \ (sum f F = 0) = (\a\F. f a = 0)" by (intro ballI sum_nonneg_eq_0_iff zero_le) context semiring_0 begin lemma sum_distrib_left: "r * sum f A = (\n\A. r * f n)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) lemma sum_distrib_right: "sum f A * r = (\n\A. f n * r)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) end lemma sum_divide_distrib: "sum f A / r = (\n\A. f n / r)" for r :: "'a::field" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (simp add: add_divide_distrib) qed lemma sum_abs[iff]: "\sum f A\ \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (auto intro: abs_triangle_ineq order_trans) qed lemma sum_abs_ge_zero[iff]: "0 \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" by (simp add: sum_nonneg) lemma abs_sum_abs[simp]: "\\a\A. \f a\\ = (\a\A. \f a\)" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then have "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp also from insert have "\ = \\f a\ + \\a\A. \f a\\\" by simp also have "\ = \f a\ + \\a\A. \f a\\" by (simp del: abs_of_nonneg) also from insert have "\ = (\a\insert a A. \f a\)" by simp finally show ?case . qed lemma sum_product: fixes f :: "'a \ 'b::semiring_0" shows "sum f A * sum g B = (\i\A. \j\B. f i * g j)" by (simp add: sum_distrib_left sum_distrib_right) (rule sum.swap) lemma sum_mult_sum_if_inj: fixes f :: "'a \ 'b::semiring_0" shows "inj_on (\(a, b). f a * g b) (A \ B) \ sum f A * sum g B = sum id {f a * g b |a b. a \ A \ b \ B}" by(auto simp: sum_product sum.cartesian_product intro!: sum.reindex_cong[symmetric]) lemma sum_SucD: "sum f A = Suc n \ \a\A. 0 < f a" by (induct A rule: infinite_finite_induct) auto lemma sum_eq_Suc0_iff: "finite A \ sum f A = Suc 0 \ (\a\A. f a = Suc 0 \ (\b\A. a \ b \ f b = 0))" by (induct A rule: finite_induct) (auto simp add: add_is_1) lemmas sum_eq_1_iff = sum_eq_Suc0_iff[simplified One_nat_def[symmetric]] lemma sum_Un_nat: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'a \ nat" \ \For the natural numbers, we have subtraction.\ by (subst sum.union_inter [symmetric]) (auto simp: algebra_simps) lemma sum_diff1_nat: "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" for f :: "'a \ nat" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case apply (auto simp: insert_Diff_if) apply (drule mk_disjoint_insert) apply auto done qed lemma sum_diff_nat: fixes f :: "'a \ nat" assumes "finite B" and "B \ A" shows "sum f (A - B) = sum f A - sum f B" using assms proof induct case empty then show ?case by simp next case (insert x F) note IH = \F \ A \ sum f (A - F) = sum f A - sum f F\ from \x \ F\ \insert x F \ A\ have "x \ A - F" by simp then have A: "sum f ((A - F) - {x}) = sum f (A - F) - f x" by (simp add: sum_diff1_nat) from \insert x F \ A\ have "F \ A" by simp with IH have "sum f (A - F) = sum f A - sum f F" by simp with A have B: "sum f ((A - F) - {x}) = sum f A - sum f F - f x" by simp from \x \ F\ have "A - insert x F = (A - F) - {x}" by auto with B have C: "sum f (A - insert x F) = sum f A - sum f F - f x" by simp from \finite F\ \x \ F\ have "sum f (insert x F) = sum f F + f x" by simp with C have "sum f (A - insert x F) = sum f A - sum f (insert x F)" by simp then show ?case by simp qed lemma sum_comp_morphism: "h 0 = 0 \ (\x y. h (x + y) = h x + h y) \ sum (h \ g) A = h (sum g A)" by (induct A rule: infinite_finite_induct) simp_all lemma (in comm_semiring_1) dvd_sum: "(\a. a \ A \ d dvd f a) \ d dvd sum f A" by (induct A rule: infinite_finite_induct) simp_all lemma (in ordered_comm_monoid_add) sum_pos: "finite I \ I \ {} \ (\i. i \ I \ 0 < f i) \ 0 < sum f I" by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos) lemma (in ordered_comm_monoid_add) sum_pos2: assumes I: "finite I" "i \ I" "0 < f i" "\i. i \ I \ 0 \ f i" shows "0 < sum f I" proof - have "0 < f i + sum f (I - {i})" using assms by (intro add_pos_nonneg sum_nonneg) auto also have "\ = sum f I" using assms by (simp add: sum.remove) finally show ?thesis . qed lemma sum_strict_mono2: fixes f :: "'a \ 'b::ordered_cancel_comm_monoid_add" assumes "finite B" "A \ B" "b \ B-A" "f b > 0" and "\x. x \ B \ f x \ 0" shows "sum f A < sum f B" proof - have "B - A \ {}" using assms(3) by blast have "sum f (B-A) > 0" by (rule sum_pos2) (use assms in auto) moreover have "sum f B = sum f (B-A) + sum f A" by (rule sum.subset_diff) (use assms in auto) ultimately show ?thesis using add_strict_increasing by auto qed lemma sum_cong_Suc: assumes "0 \ A" "\x. Suc x \ A \ f (Suc x) = g (Suc x)" shows "sum f A = sum g A" proof (rule sum.cong) fix x assume "x \ A" with assms(1) show "f x = g x" by (cases x) (auto intro!: assms(2)) qed simp_all subsubsection \Cardinality as special case of \<^const>\sum\\ lemma card_eq_sum: "card A = sum (\x. 1) A" proof - have "plus \ (\_. Suc 0) = (\_. Suc)" by (simp add: fun_eq_iff) then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" by (rule arg_cong) then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" by (blast intro: fun_cong) then show ?thesis by (simp add: card.eq_fold sum.eq_fold) qed context semiring_1 begin lemma sum_constant [simp]: "(\x \ A. y) = of_nat (card A) * y" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) +context + fixes A + assumes \finite A\ +begin + +lemma sum_of_bool_eq [simp]: + \(\x \ A. of_bool (P x)) = of_nat (card (A \ {x. P x}))\ if \finite A\ + using \finite A\ by induction simp_all + +lemma sum_mult_of_bool_eq [simp]: + \(\x \ A. f x * of_bool (P x)) = (\x \ (A \ {x. P x}). f x)\ + by (rule sum.mono_neutral_cong) (use \finite A\ in auto) + +lemma sum_of_bool_mult_eq [simp]: + \(\x \ A. of_bool (P x) * f x) = (\x \ (A \ {x. P x}). f x)\ + by (rule sum.mono_neutral_cong) (use \finite A\ in auto) + +end + end lemma sum_Suc: "sum (\x. Suc(f x)) A = sum f A + card A" using sum.distrib[of f "\_. 1" A] by simp lemma sum_bounded_above: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ f i \ K" shows "sum f A \ of_nat (card A) * K" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and g = "\x. K"] by simp next case False then show ?thesis by simp qed lemma sum_bounded_above_divide: fixes K :: "'a::linordered_field" assumes le: "\i. i\A \ f i \ K / of_nat (card A)" and fin: "finite A" "A \ {}" shows "sum f A \ K" using sum_bounded_above [of A f "K / of_nat (card A)", OF le] fin by simp lemma sum_bounded_above_strict: fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}" assumes "\i. i\A \ f i < K" "card A > 0" shows "sum f A < of_nat (card A) * K" using assms sum_strict_mono[where A=A and g = "\x. K"] by (simp add: card_gt_0_iff) lemma sum_bounded_below: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ K \ f i" shows "of_nat (card A) * K \ sum f A" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and f = "\x. K"] by simp next case False then show ?thesis by simp qed lemma convex_sum_bound_le: fixes x :: "'a \ 'b::linordered_idom" assumes 0: "\i. i \ I \ 0 \ x i" and 1: "sum x I = 1" and \: "\i. i \ I \ \a i - b\ \ \" shows "\(\i\I. a i * x i) - b\ \ \" proof - have [simp]: "(\i\I. c * x i) = c" for c by (simp flip: sum_distrib_left 1) then have "\(\i\I. a i * x i) - b\ = \\i\I. (a i - b) * x i\" by (simp add: sum_subtractf left_diff_distrib) also have "\ \ (\i\I. \(a i - b) * x i\)" using abs_abs abs_of_nonneg by blast also have "\ \ (\i\I. \(a i - b)\ * x i)" by (simp add: abs_mult 0) also have "\ \ (\i\I. \ * x i)" by (rule sum_mono) (use \ "0" mult_right_mono in blast) also have "\ = \" by simp finally show ?thesis . qed lemma card_UN_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "card (\(A ` I)) = (\i\I. card(A i))" proof - have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp with assms show ?thesis by (simp add: card_eq_sum sum.UNION_disjoint del: sum_constant) qed lemma card_Union_disjoint: assumes "pairwise disjnt C" and fin: "\A. A \ C \ finite A" shows "card (\C) = sum card C" proof (cases "finite C") case True then show ?thesis using card_UN_disjoint [OF True, of "\x. x"] assms by (simp add: disjnt_def fin pairwise_def) next case False then show ?thesis using assms card_eq_0_iff finite_UnionD by fastforce qed lemma card_Union_le_sum_card: fixes U :: "'a set set" assumes "\u \ U. finite u" shows "card (\U) \ sum card U" proof (cases "finite U") case False then show "card (\U) \ sum card U" using card_eq_0_iff finite_UnionD by auto next case True then show "card (\U) \ sum card U" proof (induct U rule: finite_induct) case empty then show ?case by auto next case (insert x F) then have "card(\(insert x F)) \ card(x) + card (\F)" using card_Un_le by auto also have "... \ card(x) + sum card F" using insert.hyps by auto also have "... = sum card (insert x F)" using sum.insert_if and insert.hyps by auto finally show ?case . qed qed lemma card_UN_le: assumes "finite I" shows "card(\i\I. A i) \ (\i\I. card(A i))" using assms proof induction case (insert i I) then show ?case using card_Un_le nat_add_left_cancel_le by (force intro: order_trans) qed auto lemma sum_multicount_gen: assumes "finite s" "finite t" "\j\t. (card {i\s. R i j} = k j)" shows "sum (\i. (card {j\t. R i j})) s = sum k t" (is "?l = ?r") proof- have "?l = sum (\i. sum (\x.1) {j\t. R i j}) s" by auto also have "\ = ?r" unfolding sum.swap_restrict [OF assms(1-2)] using assms(3) by auto finally show ?thesis . qed lemma sum_multicount: assumes "finite S" "finite T" "\j\T. (card {i\S. R i j} = k)" shows "sum (\i. card {j\T. R i j}) S = k * card T" (is "?l = ?r") proof- have "?l = sum (\i. k) T" by (rule sum_multicount_gen) (auto simp: assms) also have "\ = ?r" by (simp add: mult.commute) finally show ?thesis by auto qed lemma sum_card_image: assumes "finite A" assumes "pairwise (\s t. disjnt (f s) (f t)) A" shows "sum card (f ` A) = sum (\a. card (f a)) A" using assms proof (induct A) case (insert a A) show ?case proof cases assume "f a = {}" with insert show ?case by (subst sum.mono_neutral_right[where S="f ` A"]) (auto simp: pairwise_insert) next assume "f a \ {}" then have "sum card (insert (f a) (f ` A)) = card (f a) + sum card (f ` A)" using insert by (subst sum.insert) (auto simp: pairwise_insert) with insert show ?case by (simp add: pairwise_insert) qed qed simp subsubsection \Cardinality of products\ lemma card_SigmaI [simp]: "finite A \ \a\A. finite (B a) \ card (SIGMA x: A. B x) = (\a\A. card (B a))" by (simp add: card_eq_sum sum.Sigma del: sum_constant) (* lemma SigmaI_insert: "y \ A ==> (SIGMA x:(insert y A). B x) = (({y} \ (B y)) \ (SIGMA x: A. B x))" by auto *) lemma card_cartesian_product: "card (A \ B) = card A * card B" by (cases "finite A \ finite B") (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) lemma card_cartesian_product_singleton: "card ({x} \ A) = card A" by (simp add: card_cartesian_product) subsection \Generalized product over a set\ context comm_monoid_mult begin sublocale prod: comm_monoid_set times 1 defines prod = prod.F and prod' = prod.G .. abbreviation Prod ("\_" [1000] 999) where "\A \ prod (\x. x) A" end syntax (ASCII) "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(4PROD (_/:_)./ _)" [0, 51, 10] 10) syntax "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" == "CONST prod (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(4PROD _ |/ _./ _)" [0, 0, 10] 10) syntax "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST prod (\x. t) {x. P}" context comm_monoid_mult begin lemma prod_dvd_prod: "(\a. a \ A \ f a dvd g a) \ prod f A dvd prod g A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by (auto intro: dvdI) next case empty then show ?case by (auto intro: dvdI) next case (insert a A) then have "f a dvd g a" and "prod f A dvd prod g A" by simp_all then obtain r s where "g a = f a * r" and "prod g A = prod f A * s" by (auto elim!: dvdE) then have "g a * prod g A = f a * prod f A * (r * s)" by (simp add: ac_simps) with insert.hyps show ?case by (auto intro: dvdI) qed lemma prod_dvd_prod_subset: "finite B \ A \ B \ prod f A dvd prod f B" by (auto simp add: prod.subset_diff ac_simps intro: dvdI) end subsubsection \Properties in more restricted classes of structures\ context linordered_nonzero_semiring begin lemma prod_ge_1: "(\x. x \ A \ 1 \ f x) \ 1 \ prod f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) have "1 * 1 \ f x * prod f F" by (rule mult_mono') (use insert in auto) with insert show ?case by simp qed lemma prod_le_1: fixes f :: "'b \ 'a" assumes "\x. x \ A \ 0 \ f x \ f x \ 1" shows "prod f A \ 1" using assms proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then show ?case by (force simp: mult.commute intro: dest: mult_le_one) qed end context comm_semiring_1 begin lemma dvd_prod_eqI [intro]: assumes "finite A" and "a \ A" and "b = f a" shows "b dvd prod f A" proof - from \finite A\ have "prod f (insert a (A - {a})) = f a * prod f (A - {a})" by (intro prod.insert) auto also from \a \ A\ have "insert a (A - {a}) = A" by blast finally have "prod f A = f a * prod f (A - {a})" . with \b = f a\ show ?thesis by simp qed lemma dvd_prodI [intro]: "finite A \ a \ A \ f a dvd prod f A" by auto lemma prod_zero: assumes "finite A" and "\a\A. f a = 0" shows "prod f A = 0" using assms proof (induct A) case empty then show ?case by simp next case (insert a A) then have "f a = 0 \ (\a\A. f a = 0)" by simp then have "f a * prod f A = 0" by rule (simp_all add: insert) with insert show ?case by simp qed lemma prod_dvd_prod_subset2: assumes "finite B" and "A \ B" and "\a. a \ A \ f a dvd g a" shows "prod f A dvd prod g B" proof - from assms have "prod f A dvd prod g A" by (auto intro: prod_dvd_prod) moreover from assms have "prod g A dvd prod g B" by (auto intro: prod_dvd_prod_subset) ultimately show ?thesis by (rule dvd_trans) qed end lemma (in semidom) prod_zero_iff [simp]: fixes f :: "'b \ 'a" assumes "finite A" shows "prod f A = 0 \ (\a\A. f a = 0)" using assms by (induct A) (auto simp: no_zero_divisors) lemma (in semidom_divide) prod_diff1: assumes "finite A" and "f a \ 0" shows "prod f (A - {a}) = (if a \ A then prod f A div f a else prod f A)" proof (cases "a \ A") case True then show ?thesis by simp next case False with assms show ?thesis proof induct case empty then show ?case by simp next case (insert b B) then show ?case proof (cases "a = b") case True with insert show ?thesis by simp next case False with insert have "a \ B" by simp define C where "C = B - {a}" with \finite B\ \a \ B\ have "B = insert a C" "finite C" "a \ C" by auto with insert show ?thesis by (auto simp add: insert_commute ac_simps) qed qed qed lemma sum_zero_power [simp]: "(\i\A. c i * 0^i) = (if finite A \ 0 \ A then c 0 else 0)" for c :: "nat \ 'a::division_ring" by (induct A rule: infinite_finite_induct) auto lemma sum_zero_power' [simp]: "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" for c :: "nat \ 'a::field" using sum_zero_power [of "\i. c i / d i" A] by auto lemma (in field) prod_inversef: "prod (inverse \ f) A = inverse (prod f A)" proof (cases "finite A") case True then show ?thesis by (induct A rule: finite_induct) simp_all next case False then show ?thesis by auto qed lemma (in field) prod_dividef: "(\x\A. f x / g x) = prod f A / prod g A" using prod_inversef [of g A] by (simp add: divide_inverse prod.distrib) lemma prod_Un: fixes f :: "'b \ 'a :: field" assumes "finite A" and "finite B" and "\x\A \ B. f x \ 0" shows "prod f (A \ B) = prod f A * prod f B / prod f (A \ B)" proof - from assms have "prod f A * prod f B = prod f (A \ B) * prod f (A \ B)" by (simp add: prod.union_inter [symmetric, of A B]) with assms show ?thesis by simp qed context linordered_semidom begin lemma prod_nonneg: "(\a\A. 0 \ f a) \ 0 \ prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_pos: "(\a\A. 0 < f a) \ 0 < prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_mono: "(\i. i \ A \ 0 \ f i \ f i \ g i) \ prod f A \ prod g A" by (induct A rule: infinite_finite_induct) (force intro!: prod_nonneg mult_mono)+ lemma prod_mono_strict: assumes "finite A" "\i. i \ A \ 0 \ f i \ f i < g i" "A \ {}" shows "prod f A < prod g A" using assms proof (induct A rule: finite_induct) case empty then show ?case by simp next case insert then show ?case by (force intro: mult_strict_mono' prod_nonneg) qed end lemma prod_mono2: fixes f :: "'a \ 'b :: linordered_idom" assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 1 \ f b" and A: "\a. a \ A \ 0 \ f a" shows "prod f A \ prod f B" proof - have "prod f A \ prod f A * prod f (B-A)" by (metis prod_ge_1 A mult_le_cancel_left1 nn not_less prod_nonneg) also from fin finite_subset[OF sub fin] have "\ = prod f (A \ (B-A))" by (simp add: prod.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma less_1_prod: fixes f :: "'a \ 'b::linordered_idom" shows "finite I \ I \ {} \ (\i. i \ I \ 1 < f i) \ 1 < prod f I" by (induct I rule: finite_ne_induct) (auto intro: less_1_mult) lemma less_1_prod2: fixes f :: "'a \ 'b::linordered_idom" assumes I: "finite I" "i \ I" "1 < f i" "\i. i \ I \ 1 \ f i" shows "1 < prod f I" proof - have "1 < f i * prod f (I - {i})" using assms by (meson DiffD1 leI less_1_mult less_le_trans mult_le_cancel_left1 prod_ge_1) also have "\ = prod f I" using assms by (simp add: prod.remove) finally show ?thesis . qed lemma (in linordered_field) abs_prod: "\prod f A\ = (\x\A. \f x\)" by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult) lemma prod_eq_1_iff [simp]: "finite A \ prod f A = 1 \ (\a\A. f a = 1)" for f :: "'a \ nat" by (induct A rule: finite_induct) simp_all lemma prod_pos_nat_iff [simp]: "finite A \ prod f A > 0 \ (\a\A. f a > 0)" for f :: "'a \ nat" using prod_zero_iff by (simp del: neq0_conv add: zero_less_iff_neq_zero) lemma prod_constant [simp]: "(\x\ A. y) = y ^ card A" for y :: "'a::comm_monoid_mult" by (induct A rule: infinite_finite_induct) simp_all lemma prod_power_distrib: "prod f A ^ n = prod (\x. (f x) ^ n) A" for f :: "'a \ 'b::comm_semiring_1" by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib) lemma power_sum: "c ^ (\a\A. f a) = (\a\A. c ^ f a)" by (induct A rule: infinite_finite_induct) (simp_all add: power_add) lemma prod_gen_delta: fixes b :: "'b \ 'a::comm_monoid_mult" assumes fin: "finite S" shows "prod (\k. if k = a then b k else c) S = (if a \ S then b a * c ^ (card S - 1) else c ^ card S)" proof - let ?f = "(\k. if k=a then b k else c)" show ?thesis proof (cases "a \ S") case False then have "\ k\ S. ?f k = c" by simp with False show ?thesis by (simp add: prod_constant) next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have disjoint: "?A \ ?B = {}" by simp from fin have fin': "finite ?A" "finite ?B" by auto have f_A0: "prod ?f ?A = prod (\i. c) ?A" by (rule prod.cong) auto from fin True have card_A: "card ?A = card S - 1" by auto have f_A1: "prod ?f ?A = c ^ card ?A" unfolding f_A0 by (rule prod_constant) have "prod ?f ?A * prod ?f ?B = prod ?f S" using prod.union_disjoint[OF fin' disjoint, of ?f, unfolded eq[symmetric]] by simp with True card_A show ?thesis by (simp add: f_A1 field_simps cong add: prod.cong cong del: if_weak_cong) qed qed lemma sum_image_le: fixes g :: "'a \ 'b::ordered_comm_monoid_add" assumes "finite I" "\i. i \ I \ 0 \ g(f i)" shows "sum g (f ` I) \ sum (g \ f) I" using assms proof induction case empty then show ?case by auto next case (insert x F) from insertI1 have "0 \ g (f x)" by (rule insert) hence 1: "sum g (f ` F) \ g (f x) + sum g (f ` F)" using add_increasing by blast have 2: "sum g (f ` F) \ sum (g \ f) F" using insert by blast have "sum g (f ` insert x F) = sum g (insert (f x) (f ` F))" by simp also have "\ \ g (f x) + sum g (f ` F)" by (simp add: 1 insert sum.insert_if) also from 2 have "\ \ g (f x) + sum (g \ f) F" by (rule add_left_mono) also from insert(1, 2) have "\ = sum (g \ f) (insert x F)" by (simp add: sum.insert_if) finally show ?case . qed end diff --git a/src/HOL/Groups_List.thy b/src/HOL/Groups_List.thy --- a/src/HOL/Groups_List.thy +++ b/src/HOL/Groups_List.thy @@ -1,582 +1,583 @@ (* Author: Tobias Nipkow, TU Muenchen *) section \Sum and product over lists\ theory Groups_List imports List begin locale monoid_list = monoid begin definition F :: "'a list \ 'a" where eq_foldr [code]: "F xs = foldr f xs \<^bold>1" lemma Nil [simp]: "F [] = \<^bold>1" by (simp add: eq_foldr) lemma Cons [simp]: "F (x # xs) = x \<^bold>* F xs" by (simp add: eq_foldr) lemma append [simp]: "F (xs @ ys) = F xs \<^bold>* F ys" by (induct xs) (simp_all add: assoc) end locale comm_monoid_list = comm_monoid + monoid_list begin lemma rev [simp]: "F (rev xs) = F xs" by (simp add: eq_foldr foldr_fold fold_rev fun_eq_iff assoc left_commute) end locale comm_monoid_list_set = list: comm_monoid_list + set: comm_monoid_set begin lemma distinct_set_conv_list: "distinct xs \ set.F g (set xs) = list.F (map g xs)" by (induct xs) simp_all lemma set_conv_list [code]: "set.F g (set xs) = list.F (map g (remdups xs))" by (simp add: distinct_set_conv_list [symmetric]) end subsection \List summation\ context monoid_add begin sublocale sum_list: monoid_list plus 0 defines sum_list = sum_list.F .. end context comm_monoid_add begin sublocale sum_list: comm_monoid_list plus 0 rewrites "monoid_list.F plus 0 = sum_list" proof - show "comm_monoid_list plus 0" .. then interpret sum_list: comm_monoid_list plus 0 . from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp qed sublocale sum: comm_monoid_list_set plus 0 rewrites "monoid_list.F plus 0 = sum_list" and "comm_monoid_set.F plus 0 = sum" proof - show "comm_monoid_list_set plus 0" .. then interpret sum: comm_monoid_list_set plus 0 . from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym) qed end text \Some syntactic sugar for summing a function over a list:\ syntax (ASCII) "_sum_list" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10) syntax "_sum_list" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\x\xs. b" == "CONST sum_list (CONST map (\x. b) xs)" context includes lifting_syntax begin lemma sum_list_transfer [transfer_rule]: "(list_all2 A ===> A) sum_list sum_list" if [transfer_rule]: "A 0 0" "(A ===> A ===> A) (+) (+)" unfolding sum_list.eq_foldr [abs_def] by transfer_prover end text \TODO duplicates\ lemmas sum_list_simps = sum_list.Nil sum_list.Cons lemmas sum_list_append = sum_list.append lemmas sum_list_rev = sum_list.rev lemma (in monoid_add) fold_plus_sum_list_rev: "fold plus xs = plus (sum_list (rev xs))" proof fix x have "fold plus xs x = sum_list (rev xs @ [x])" by (simp add: foldr_conv_fold sum_list.eq_foldr) also have "\ = sum_list (rev xs) + x" by simp finally show "fold plus xs x = sum_list (rev xs) + x" . qed lemma (in comm_monoid_add) sum_list_map_remove1: "x \ set xs \ sum_list (map f xs) = f x + sum_list (map f (remove1 x xs))" by (induct xs) (auto simp add: ac_simps) lemma (in monoid_add) size_list_conv_sum_list: "size_list f xs = sum_list (map f xs) + size xs" by (induct xs) auto lemma (in monoid_add) length_concat: "length (concat xss) = sum_list (map length xss)" by (induct xss) simp_all lemma (in monoid_add) length_product_lists: "length (product_lists xss) = foldr (*) (map length xss) 1" proof (induct xss) case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def) qed simp lemma (in monoid_add) sum_list_map_filter: assumes "\x. x \ set xs \ \ P x \ f x = 0" shows "sum_list (map f (filter P xs)) = sum_list (map f xs)" using assms by (induct xs) auto lemma sum_list_filter_le_nat: fixes f :: "'a \ nat" shows "sum_list (map f (filter P xs)) \ sum_list (map f xs)" by(induction xs; simp) lemma (in comm_monoid_add) distinct_sum_list_conv_Sum: "distinct xs \ sum_list xs = Sum (set xs)" by (induct xs) simp_all lemma sum_list_upt[simp]: "m \ n \ sum_list [m.. {m..x. x \ set xs \ 0 \ x) \ 0 \ sum_list xs" by (induction xs) auto lemma sum_list_nonpos: "(\x. x \ set xs \ x \ 0) \ sum_list xs \ 0" by (induction xs) (auto simp: add_nonpos_nonpos) lemma sum_list_nonneg_eq_0_iff: "(\x. x \ set xs \ 0 \ x) \ sum_list xs = 0 \ (\x\ set xs. x = 0)" by (induction xs) (simp_all add: add_nonneg_eq_0_iff sum_list_nonneg) end context canonically_ordered_monoid_add begin lemma sum_list_eq_0_iff [simp]: "sum_list ns = 0 \ (\n \ set ns. n = 0)" by (simp add: sum_list_nonneg_eq_0_iff) lemma member_le_sum_list: "x \ set xs \ x \ sum_list xs" by (induction xs) (auto simp: add_increasing add_increasing2) lemma elem_le_sum_list: "k < size ns \ ns ! k \ sum_list (ns)" by (rule member_le_sum_list) simp end lemma (in ordered_cancel_comm_monoid_diff) sum_list_update: "k < size xs \ sum_list (xs[k := x]) = sum_list xs + x - xs ! k" apply(induction xs arbitrary:k) apply (auto simp: add_ac split: nat.split) apply(drule elem_le_sum_list) by (simp add: local.add_diff_assoc local.add_increasing) lemma (in monoid_add) sum_list_triv: "(\x\xs. r) = of_nat (length xs) * r" by (induct xs) (simp_all add: distrib_right) lemma (in monoid_add) sum_list_0 [simp]: "(\x\xs. 0) = 0" by (induct xs) (simp_all add: distrib_right) text\For non-Abelian groups \xs\ needs to be reversed on one side:\ lemma (in ab_group_add) uminus_sum_list_map: "- sum_list (map f xs) = sum_list (map (uminus \ f) xs)" by (induct xs) simp_all lemma (in comm_monoid_add) sum_list_addf: "(\x\xs. f x + g x) = sum_list (map f xs) + sum_list (map g xs)" by (induct xs) (simp_all add: algebra_simps) lemma (in ab_group_add) sum_list_subtractf: "(\x\xs. f x - g x) = sum_list (map f xs) - sum_list (map g xs)" by (induct xs) (simp_all add: algebra_simps) lemma (in semiring_0) sum_list_const_mult: "(\x\xs. c * f x) = c * (\x\xs. f x)" by (induct xs) (simp_all add: algebra_simps) lemma (in semiring_0) sum_list_mult_const: "(\x\xs. f x * c) = (\x\xs. f x) * c" by (induct xs) (simp_all add: algebra_simps) lemma (in ordered_ab_group_add_abs) sum_list_abs: "\sum_list xs\ \ sum_list (map abs xs)" by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq]) lemma sum_list_mono: fixes f g :: "'a \ 'b::{monoid_add, ordered_ab_semigroup_add}" shows "(\x. x \ set xs \ f x \ g x) \ (\x\xs. f x) \ (\x\xs. g x)" by (induct xs) (simp, simp add: add_mono) lemma sum_list_strict_mono: fixes f g :: "'a \ 'b::{monoid_add, strict_ordered_ab_semigroup_add}" shows "\ xs \ []; \x. x \ set xs \ f x < g x \ \ sum_list (map f xs) < sum_list (map g xs)" proof (induction xs) case Nil thus ?case by simp next case C: (Cons _ xs) show ?case proof (cases xs) case Nil thus ?thesis using C.prems by simp next case Cons thus ?thesis using C by(simp add: add_strict_mono) qed qed lemma (in monoid_add) sum_list_distinct_conv_sum_set: "distinct xs \ sum_list (map f xs) = sum f (set xs)" by (induct xs) simp_all lemma (in monoid_add) interv_sum_list_conv_sum_set_nat: "sum_list (map f [m..General equivalence between \<^const>\sum_list\ and \<^const>\sum\\ lemma (in monoid_add) sum_list_sum_nth: "sum_list xs = (\ i = 0 ..< length xs. xs ! i)" using interv_sum_list_conv_sum_set_nat [of "(!) xs" 0 "length xs"] by (simp add: map_nth) lemma sum_list_map_eq_sum_count: "sum_list (map f xs) = sum (\x. count_list xs x * f x) (set xs)" proof(induction xs) case (Cons x xs) show ?case (is "?l = ?r") proof cases assume "x \ set xs" have "?l = f x + (\x\set xs. count_list xs x * f x)" by (simp add: Cons.IH) also have "set xs = insert x (set xs - {x})" using \x \ set xs\by blast also have "f x + (\x\insert x (set xs - {x}). count_list xs x * f x) = ?r" by (simp add: sum.insert_remove eq_commute) finally show ?thesis . next assume "x \ set xs" hence "\xa. xa \ set xs \ x \ xa" by blast thus ?thesis by (simp add: Cons.IH \x \ set xs\) qed qed simp lemma sum_list_map_eq_sum_count2: assumes "set xs \ X" "finite X" shows "sum_list (map f xs) = sum (\x. count_list xs x * f x) X" proof- let ?F = "\x. count_list xs x * f x" have "sum ?F X = sum ?F (set xs \ (X - set xs))" using Un_absorb1[OF assms(1)] by(simp) also have "\ = sum ?F (set xs)" using assms(2) by(simp add: sum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel) finally show ?thesis by(simp add:sum_list_map_eq_sum_count) qed lemma sum_list_replicate: "sum_list (replicate n c) = of_nat n * c" by(induction n)(auto simp add: distrib_right) lemma sum_list_nonneg: "(\x. x \ set xs \ (x :: 'a :: ordered_comm_monoid_add) \ 0) \ sum_list xs \ 0" by (induction xs) simp_all lemma sum_list_Suc: "sum_list (map (\x. Suc(f x)) xs) = sum_list (map f xs) + length xs" by(induction xs; simp) lemma (in monoid_add) sum_list_map_filter': "sum_list (map f (filter P xs)) = sum_list (map (\x. if P x then f x else 0) xs)" by (induction xs) simp_all text \Summation of a strictly ascending sequence with length \n\ can be upper-bounded by summation over \{0...\ lemma sorted_wrt_less_sum_mono_lowerbound: fixes f :: "nat \ ('b::ordered_comm_monoid_add)" assumes mono: "\x y. x\y \ f x \ f y" shows "sorted_wrt (<) ns \ (\i\{0.. (\i\ns. f i)" proof (induction ns rule: rev_induct) case Nil then show ?case by simp next case (snoc n ns) have "sum f {0.. sum_list (map f ns)" using snoc by (auto simp: sorted_wrt_append) also have "length ns \ n" using sorted_wrt_less_idx[OF snoc.prems(1), of "length ns"] by auto finally have "sum f {0.. sum_list (map f ns) + f n" using mono add_mono by blast thus ?case by simp qed subsection \Horner sums\ context comm_semiring_0 begin definition horner_sum :: \('b \ 'a) \ 'a \ 'b list \ 'a\ where horner_sum_foldr: \horner_sum f a xs = foldr (\x b. f x + a * b) xs 0\ lemma horner_sum_simps [simp]: \horner_sum f a [] = 0\ \horner_sum f a (x # xs) = f x + a * horner_sum f a xs\ by (simp_all add: horner_sum_foldr) lemma horner_sum_eq_sum_funpow: \horner_sum f a xs = (\n = 0.. proof (induction xs) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (simp add: sum.atLeast0_lessThan_Suc_shift sum_distrib_left del: sum.op_ivl_Suc) qed end context includes lifting_syntax begin lemma horner_sum_transfer [transfer_rule]: \((B ===> A) ===> A ===> list_all2 B ===> A) horner_sum horner_sum\ if [transfer_rule]: \A 0 0\ and [transfer_rule]: \(A ===> A ===> A) (+) (+)\ and [transfer_rule]: \(A ===> A ===> A) (*) (*)\ by (unfold horner_sum_foldr) transfer_prover end context comm_semiring_1 begin lemma horner_sum_eq_sum: \horner_sum f a xs = (\n = 0.. proof - have \(*) a ^^ n = (*) (a ^ n)\ for n by (induction n) (simp_all add: ac_simps) then show ?thesis by (simp add: horner_sum_eq_sum_funpow ac_simps) qed lemma horner_sum_append: \horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\ using sum.atLeastLessThan_shift_bounds [of _ 0 \length xs\ \length ys\] atLeastLessThan_add_Un [of 0 \length xs\ \length ys\] by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add) end context semiring_bit_shifts begin lemma horner_sum_bit_eq_take_bit: \horner_sum of_bool 2 (map (bit a) [0.. proof (induction a arbitrary: n rule: bits_induct) case (stable a) moreover have \bit a = (\_. odd a)\ using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff) moreover have \{q. q < n} = {0.. by auto ultimately show ?case by (simp add: stable_imp_take_bit_eq horner_sum_eq_sum mask_eq_sum_exp) next case (rec a b) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc m) have \map (bit (of_bool b + 2 * a)) [0.. by (simp only: upt_conv_Cons) simp also have \\ = b # map (bit a) [0.. by (simp only: flip: map_Suc_upt) (simp add: bit_Suc rec.hyps) finally show ?thesis - using Suc rec.IH [of m] by (simp add: take_bit_Suc rec.hyps, simp add: ac_simps mod_2_eq_odd) + using Suc rec.IH [of m] by (simp add: take_bit_Suc rec.hyps) + (simp_all add: ac_simps mod_2_eq_odd) qed qed end context unique_euclidean_semiring_with_bit_shifts begin lemma bit_horner_sum_bit_iff [bit_simps]: \bit (horner_sum of_bool 2 bs) n \ n < length bs \ bs ! n\ 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 m) with bit_rec [of _ n] Cons.prems Cons.IH [of m] show ?thesis by simp qed qed lemma take_bit_horner_sum_bit_eq: \take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff) end lemma horner_sum_of_bool_2_less: \(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\ proof - have \(\n = 0.. (\n = 0.. by (rule sum_mono) simp also have \\ = 2 ^ length bs - 1\ by (induction bs) simp_all finally show ?thesis by (simp add: horner_sum_eq_sum) qed subsection \Further facts about \<^const>\List.n_lists\\ lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n" by (induct n) (auto simp add: comp_def length_concat sum_list_triv) lemma distinct_n_lists: assumes "distinct xs" shows "distinct (List.n_lists n xs)" proof (rule card_distinct) from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) have "card (set (List.n_lists n xs)) = card (set xs) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) moreover have "card (\ys\set (List.n_lists n xs). (\y. y # ys) ` set xs) = (\ys\set (List.n_lists n xs). card ((\y. y # ys) ` set xs))" by (rule card_UN_disjoint) auto moreover have "\ys. card ((\y. y # ys) ` set xs) = card (set xs)" by (rule card_image) (simp add: inj_on_def) ultimately show ?case by auto qed also have "\ = length xs ^ n" by (simp add: card_length) finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)" by (simp add: length_n_lists) qed subsection \Tools setup\ lemmas sum_code = sum.set_conv_list lemma sum_set_upto_conv_sum_list_int [code_unfold]: "sum f (set [i..j::int]) = sum_list (map f [i..j])" by (simp add: interv_sum_list_conv_sum_set_int) lemma sum_set_upt_conv_sum_list_nat [code_unfold]: "sum f (set [m..List product\ context monoid_mult begin sublocale prod_list: monoid_list times 1 defines prod_list = prod_list.F .. end context comm_monoid_mult begin sublocale prod_list: comm_monoid_list times 1 rewrites "monoid_list.F times 1 = prod_list" proof - show "comm_monoid_list times 1" .. then interpret prod_list: comm_monoid_list times 1 . from prod_list_def show "monoid_list.F times 1 = prod_list" by simp qed sublocale prod: comm_monoid_list_set times 1 rewrites "monoid_list.F times 1 = prod_list" and "comm_monoid_set.F times 1 = prod" proof - show "comm_monoid_list_set times 1" .. then interpret prod: comm_monoid_list_set times 1 . from prod_list_def show "monoid_list.F times 1 = prod_list" by simp from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym) qed end text \Some syntactic sugar:\ syntax (ASCII) "_prod_list" :: "pttrn => 'a list => 'b => 'b" ("(3PROD _<-_. _)" [0, 51, 10] 10) syntax "_prod_list" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\x\xs. b" \ "CONST prod_list (CONST map (\x. b) xs)" context includes lifting_syntax begin lemma prod_list_transfer [transfer_rule]: "(list_all2 A ===> A) prod_list prod_list" if [transfer_rule]: "A 1 1" "(A ===> A ===> A) (*) (*)" unfolding prod_list.eq_foldr [abs_def] by transfer_prover end lemma prod_list_zero_iff: "prod_list xs = 0 \ (0 :: 'a :: {semiring_no_zero_divisors, semiring_1}) \ set xs" by (induction xs) simp_all 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,1850 +1,1855 @@ (* Author: Florian Haftmann, TUM *) section \Bit operations in suitable algebraic structures\ theory Bit_Operations imports Main "HOL-Library.Boolean_Algebra" 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 [bit_simps]: \\n. bit (a AND b) n \ bit a n \ bit b n\ and bit_or_iff [bit_simps]: \\n. bit (a OR b) n \ bit a n \ bit b n\ and bit_xor_iff [bit_simps]: \\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_simps]: \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_mask [simp]: \take_bit m (mask n) = mask (min m n)\ by (rule bit_eqI) (simp add: bit_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 or_eq_0_iff: \a OR b = 0 \ a = 0 \ b = 0\ by (auto simp add: bit_eq_iff bit_or_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\) lemma bit_iff_and_drop_bit_eq_1: \bit a n \ drop_bit n a AND 1 = 1\ by (simp add: bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one) lemma bit_iff_and_push_bit_not_eq_0: \bit a n \ a AND push_bit n 1 \ 0\ apply (cases \2 ^ n = 0\) apply (simp_all add: push_bit_of_1 bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit) apply (simp_all add: bit_exp_iff) done end class ring_bit_operations = semiring_bit_operations + ring_parity + fixes not :: \'a \ 'a\ (\NOT\) assumes bit_not_iff [bit_simps]: \\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_simps]: \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_simps]: \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_simps]: \bit (- (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1) 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 (in ring_bit_operations) and_eq_minus_1_iff: \a AND b = - 1 \ a = - 1 \ b = - 1\ proof assume \a = - 1 \ b = - 1\ then show \a AND b = - 1\ by simp next assume \a AND b = - 1\ have *: \bit a n\ \bit b n\ if \2 ^ n \ 0\ for n proof - from \a AND b = - 1\ have \bit (a AND b) n = bit (- 1) n\ by (simp add: bit_eq_iff) then show \bit a n\ \bit b n\ using that by (simp_all add: bit_and_iff) qed have \a = - 1\ by (rule bit_eqI) (simp add: *) moreover have \b = - 1\ by (rule bit_eqI) (simp add: *) ultimately show \a = - 1 \ b = - 1\ by simp qed 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_simps]: \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_simps]: \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_simps]: \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\\ lemma int_bit_bound: fixes k :: int obtains n where \\m. n \ m \ bit k m \ bit k n\ and \n > 0 \ bit k (n - 1) \ bit k n\ proof - obtain q where *: \\m. q \ m \ bit k m \ bit k q\ proof (cases \k \ 0\) case True moreover from power_gt_expt [of 2 \nat k\] have \k < 2 ^ nat k\ by simp ultimately have *: \k div 2 ^ nat k = 0\ by simp show thesis proof (rule that [of \nat k\]) fix m assume \nat k \ m\ then show \bit k m \ bit k (nat k)\ by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex) qed next case False moreover from power_gt_expt [of 2 \nat (- k)\] have \- k \ 2 ^ nat (- k)\ by simp ultimately have \- k div - (2 ^ nat (- k)) = - 1\ by (subst div_pos_neg_trivial) simp_all then have *: \k div 2 ^ nat (- k) = - 1\ by simp show thesis proof (rule that [of \nat (- k)\]) fix m assume \nat (- k) \ m\ then show \bit k m \ bit k (nat (- k))\ by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex) qed qed show thesis proof (cases \\m. bit k m \ bit k q\) case True then have \bit k 0 \ bit k q\ by blast with True that [of 0] show thesis by simp next case False then obtain r where **: \bit k r \ bit k q\ by blast have \r < q\ by (rule ccontr) (use * [of r] ** in simp) define N where \N = {n. n < q \ bit k n \ bit k q}\ moreover have \finite N\ \r \ N\ using ** N_def \r < q\ by auto moreover define n where \n = Suc (Max N)\ ultimately have \\m. n \ m \ bit k m \ bit k n\ apply auto apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \finite N\ all_not_in_conv mem_Collect_eq not_le) apply (metis "*" Max_ge Suc_n_not_le_n \finite N\ linorder_not_less mem_Collect_eq) apply (metis "*" Max_ge Suc_n_not_le_n \finite N\ linorder_not_less mem_Collect_eq) apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \finite N\ all_not_in_conv mem_Collect_eq not_le) done have \bit k (Max N) \ bit k n\ by (metis (mono_tags, lifting) "*" Max_in N_def \\m. n \ m \ bit k m = bit k n\ \finite N\ \r \ N\ empty_iff le_cases mem_Collect_eq) show thesis apply (rule that [of n]) using \\m. n \ m \ bit k m = bit k n\ apply blast using \bit k (Max N) \ bit k n\ n_def by auto qed qed 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_simps]: \bit (NOT k) n \ \ bit k n\ for k :: int by (simp add: bit_not_int_iff' not_int_def) 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) + (simp_all 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 OR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x OR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) qed lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x XOR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) qed lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "0 \ x AND y" using assms by simp lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x OR y" using assms by simp lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x XOR y" using assms by simp lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "x AND y \ x" - using assms by (induction x arbitrary: y rule: int_bit_induct) - (simp_all add: and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] add_increasing) +using assms proof (induction x arbitrary: y rule: int_bit_induct) + case (odd k) + then have \k AND y div 2 \ k\ + by simp + then show ?case + by (simp add: and_int_rec [of \1 + _ * 2\]) +qed (simp_all add: and_int_rec [of \_ * 2\]) lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ y" shows "x AND y \ y" using assms AND_upper1 [of y x] by (simp add: ac_simps) lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemma plus_and_or: \(x AND y) + (x OR y) = x + y\ for x y :: int proof (induction x arbitrary: y rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) next case (odd x) from odd.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) qed 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 lemma take_bit_eq_mask_iff: \take_bit n k = mask n \ take_bit n (k + 1) = 0\ (is \?P \ ?Q\) for k :: int proof assume ?P then have \take_bit n (take_bit n k + take_bit n 1) = 0\ by (simp add: mask_eq_exp_minus_1) then show ?Q by (simp only: take_bit_add) next assume ?Q then have \take_bit n (k + 1) - 1 = - 1\ by simp then have \take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\ by simp moreover have \take_bit n (take_bit n (k + 1) - 1) = take_bit n k\ by (simp add: take_bit_eq_mod mod_simps) ultimately show ?P by (simp add: take_bit_minus_one_eq_mask) qed lemma take_bit_eq_mask_iff_exp_dvd: \take_bit n k = mask n \ 2 ^ n dvd k + 1\ for k :: int by (simp add: take_bit_eq_mask_iff flip: take_bit_eq_0_iff) 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_simps]: \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 text \FIXME: The rule sets below are very large (24 rules for each operator). Is there a simpler way to do this?\ context begin private lemma eqI: \k = l\ if num: \\n. bit k (numeral n) \ bit l (numeral n)\ and even: \even k \ even l\ for k l :: int proof (rule bit_eqI) fix n show \bit k n \ bit l n\ proof (cases n) case 0 with even show ?thesis by simp next case (Suc n) with num [of \num_of_nat (Suc n)\] show ?thesis by (simp only: numeral_num_of_nat) qed qed lemma int_and_numerals [simp]: "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))" "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))" "(1::int) AND numeral (Num.Bit0 y) = 0" "(1::int) AND numeral (Num.Bit1 y) = 1" "(1::int) AND - numeral (Num.Bit0 y) = 0" "(1::int) AND - numeral (Num.Bit1 y) = 1" "numeral (Num.Bit0 x) AND (1::int) = 0" "numeral (Num.Bit1 x) AND (1::int) = 1" "- numeral (Num.Bit0 x) AND (1::int) = 0" "- numeral (Num.Bit1 x) AND (1::int) = 1" by (auto simp add: bit_and_iff bit_minus_iff even_and_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq intro: eqI) lemma int_or_numerals [simp]: "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))" "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" by (auto simp add: bit_or_iff bit_minus_iff even_or_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) lemma int_xor_numerals [simp]: "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))" "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" by (auto simp add: bit_xor_iff bit_minus_iff even_xor_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) 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_simps]: \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) lemma concat_bit_take_bit_eq: \concat_bit n (take_bit n b) = concat_bit n b\ by (simp add: concat_bit_def [abs_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_simps]: \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_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_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 lemma not_minus_numeral_inc_eq: \NOT (- numeral (Num.inc n)) = (numeral n :: int)\ by (simp add: not_int_def sub_inc_One_eq) 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 (simp add: and_nat_def bit_simps) show \bit (m OR n) q \ bit m q \ bit n q\ by (simp add: or_nat_def bit_simps) show \bit (m XOR n) q \ bit m q \ bit n q\ by (simp add: xor_nat_def bit_simps) 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 lemma Suc_mask_eq_exp: \Suc (mask n) = 2 ^ n\ by (simp add: mask_eq_exp_minus_1) lemma less_eq_mask: \n \ mask n\ by (simp add: mask_eq_exp_minus_1 le_diff_conv2) (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0) lemma less_mask: \n < mask n\ if \Suc 0 < n\ proof - define m where \m = n - 2\ with that have *: \n = m + 2\ by simp have \Suc (Suc (Suc m)) < 4 * 2 ^ m\ by (induction m) simp_all then have \Suc (m + 2) < Suc (mask (m + 2))\ by (simp add: Suc_mask_eq_exp) then have \m + 2 < mask (m + 2)\ by (simp add: less_le) with * show ?thesis by simp qed 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]} \ code_identifier type_class semiring_bits \ (SML) Bit_Operations.semiring_bits and (OCaml) Bit_Operations.semiring_bits and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bits | class_relation semiring_bits < semiring_parity \ (SML) Bit_Operations.semiring_parity_semiring_bits and (OCaml) Bit_Operations.semiring_parity_semiring_bits and (Haskell) Bit_Operations.semiring_parity_semiring_bits and (Scala) Bit_Operations.semiring_parity_semiring_bits | constant bit \ (SML) Bit_Operations.bit and (OCaml) Bit_Operations.bit and (Haskell) Bit_Operations.bit and (Scala) Bit_Operations.bit | class_instance nat :: semiring_bits \ (SML) Bit_Operations.semiring_bits_nat and (OCaml) Bit_Operations.semiring_bits_nat and (Haskell) Bit_Operations.semiring_bits_nat and (Scala) Bit_Operations.semiring_bits_nat | class_instance int :: semiring_bits \ (SML) Bit_Operations.semiring_bits_int and (OCaml) Bit_Operations.semiring_bits_int and (Haskell) Bit_Operations.semiring_bits_int and (Scala) Bit_Operations.semiring_bits_int | type_class semiring_bit_shifts \ (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bit_shifts | class_relation semiring_bit_shifts < semiring_bits \ (SML) Bit_Operations.semiring_bits_semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bits_semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits_semiring_bit_shifts and (Scala) Bit_Operations.semiring_bits_semiring_bit_shifts | constant push_bit \ (SML) Bit_Operations.push_bit and (OCaml) Bit_Operations.push_bit and (Haskell) Bit_Operations.push_bit and (Scala) Bit_Operations.push_bit | constant drop_bit \ (SML) Bit_Operations.drop_bit and (OCaml) Bit_Operations.drop_bit and (Haskell) Bit_Operations.drop_bit and (Scala) Bit_Operations.drop_bit | constant take_bit \ (SML) Bit_Operations.take_bit and (OCaml) Bit_Operations.take_bit and (Haskell) Bit_Operations.take_bit and (Scala) Bit_Operations.take_bit | class_instance nat :: semiring_bit_shifts \ (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts | class_instance int :: semiring_bit_shifts \ (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts end diff --git a/src/HOL/Library/Word.thy b/src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy +++ b/src/HOL/Library/Word.thy @@ -1,4368 +1,4368 @@ (* Title: HOL/Library/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA, et. al. *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" begin subsection \Preliminaries\ lemma signed_take_bit_decr_length_iff: \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by (cases \LENGTH('a)\) (simp_all add: signed_take_bit_eq_iff_take_bit_eq) subsection \Fundamentals\ subsubsection \Type definition\ quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) hide_const (open) rep \ \only for foundational purpose\ hide_const (open) Word \ \only for code generation\ subsubsection \Basic arithmetic\ instantiation word :: (len) comm_ring_1 begin lift_definition zero_word :: \'a word\ is 0 . lift_definition one_word :: \'a word\ is 1 . lift_definition plus_word :: \'a word \ 'a word \ 'a word\ is \(+)\ by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition minus_word :: \'a word \ 'a word \ 'a word\ is \(-)\ by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lift_definition uminus_word :: \'a word \ 'a word\ is uminus by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) lift_definition times_word :: \'a word \ 'a word \ 'a word\ is \(*)\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by (standard; transfer) (simp_all add: algebra_simps) end context includes lifting_syntax notes power_transfer [transfer_rule] transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] transfer_rule_of_nat [transfer_rule] transfer_rule_of_int [transfer_rule] begin lemma power_transfer_word [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) of_bool of_bool\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) numeral numeral\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) int of_nat\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) (\k. k) of_int\ proof - have \((=) ===> pcr_word) of_int of_int\ by transfer_prover then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed end lemma exp_eq_zero_iff [simp]: \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ - by transfer simp + by transfer auto lemma word_exp_length_eq_0 [simp]: \(2 :: 'a::len word) ^ LENGTH('a) = 0\ by simp subsubsection \Basic tool setup\ ML_file \Tools/word_lib.ML\ subsubsection \Basic code generation setup\ context begin qualified lift_definition the_int :: \'a::len word \ int\ is \take_bit LENGTH('a)\ . end lemma [code abstype]: \Word.Word (Word.the_int w) = w\ by transfer simp lemma Word_eq_word_of_int [code_post, simp]: \Word.Word = of_int\ by (rule; transfer) simp quickcheck_generator word constructors: \0 :: 'a::len word\, \numeral :: num \ 'a::len word\ instantiation word :: (len) equal begin lift_definition equal_word :: \'a word \ 'a word \ bool\ is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by simp instance by (standard; transfer) rule end lemma [code]: \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ by transfer (simp add: equal) lemma [code]: \Word.the_int 0 = 0\ by transfer simp lemma [code]: \Word.the_int 1 = 1\ by transfer simp lemma [code]: \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_add) lemma [code]: \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ for w :: \'a::len word\ by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) lemma [code]: \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_diff) lemma [code]: \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_mult) subsubsection \Basic conversions\ abbreviation word_of_nat :: \nat \ 'a::len word\ where \word_of_nat \ of_nat\ abbreviation word_of_int :: \int \ 'a::len word\ where \word_of_int \ of_int\ lemma word_of_nat_eq_iff: \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_eq_iff: \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_eq_0_iff [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 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) lemma inj_unsigned [simp]: \inj unsigned\ by (rule injI) (simp add: unsigned_word_eqI) lemma unsigned_eq_0_iff: \unsigned w = 0 \ w = 0\ using word_eq_iff_unsigned [of w 0] by simp end context ring_1 begin lift_definition signed :: \'b::len word \ 'a\ is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ by (simp flip: signed_take_bit_decr_length_iff) lemma signed_0 [simp]: \signed 0 = 0\ by transfer simp lemma signed_1 [simp]: \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ by (transfer fixing: uminus; cases \LENGTH('b)\) (auto dest: gr0_implies_Suc) lemma signed_minus_1 [simp]: \signed (- 1 :: 'b::len word) = - 1\ by (transfer fixing: uminus) simp lemma signed_numeral [simp]: \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ by transfer simp lemma signed_neg_numeral [simp]: \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ by transfer simp lemma signed_of_nat [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) lemma inj_signed [simp]: \inj signed\ by (rule injI) (simp add: signed_word_eqI) lemma signed_eq_0_iff: \signed w = 0 \ w = 0\ using word_eq_iff_signed [of w 0] by simp end abbreviation unat :: \'a::len word \ nat\ where \unat \ unsigned\ abbreviation uint :: \'a::len word \ int\ where \uint \ unsigned\ abbreviation sint :: \'a::len word \ int\ where \sint \ signed\ abbreviation ucast :: \'a::len word \ 'b::len word\ where \ucast \ unsigned\ abbreviation scast :: \'a::len word \ 'b::len word\ where \scast \ signed\ context includes lifting_syntax begin lemma [transfer_rule]: \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ using unsigned.transfer [where ?'a = nat] by simp lemma [transfer_rule]: \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ using unsigned.transfer [where ?'a = int] by (simp add: comp_def) lemma [transfer_rule]: \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ using signed.transfer [where ?'a = int] by simp lemma [transfer_rule]: \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ by simp qed lemma [transfer_rule]: \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ by simp qed end lemma of_nat_unat [simp]: \of_nat (unat w) = unsigned w\ by transfer simp lemma of_int_uint [simp]: \of_int (uint w) = unsigned w\ by transfer simp lemma of_int_sint [simp]: \of_int (sint a) = signed a\ by transfer (simp_all add: take_bit_signed_take_bit) lemma nat_uint_eq [simp]: \nat (uint w) = unat w\ by transfer simp lemma sgn_uint_eq [simp]: \sgn (uint w) = of_bool (w \ 0)\ by transfer (simp add: less_le) text \Aliasses only for code generation\ context begin qualified lift_definition of_int :: \int \ 'a::len word\ is \take_bit LENGTH('a)\ . qualified lift_definition of_nat :: \nat \ 'a::len word\ is \int \ take_bit LENGTH('a)\ . qualified lift_definition the_nat :: \'a::len word \ nat\ is \nat \ take_bit LENGTH('a)\ by simp qualified lift_definition the_signed_int :: \'a::len word \ int\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) qualified lift_definition cast :: \'a::len word \ 'b::len word\ is \take_bit LENGTH('a)\ by simp qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) end lemma [code_abbrev, simp]: \Word.the_int = uint\ by transfer rule lemma [code]: \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer simp lemma [code_abbrev, simp]: \Word.of_int = word_of_int\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ by transfer (simp add: take_bit_of_nat) lemma [code_abbrev, simp]: \Word.of_nat = word_of_nat\ by (rule; transfer) (simp add: take_bit_of_nat) lemma [code]: \Word.the_nat w = nat (Word.the_int w)\ by transfer simp lemma [code_abbrev, simp]: \Word.the_nat = unat\ by (rule; transfer) simp lemma [code]: \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: signed_take_bit_take_bit) lemma [code_abbrev, simp]: \Word.the_signed_int = sint\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.cast = ucast\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.signed_cast = scast\ by (rule; transfer) simp lemma [code]: \unsigned w = of_nat (nat (Word.the_int w))\ by transfer simp lemma [code]: \signed w = of_int (Word.the_signed_int w)\ by transfer simp subsubsection \Basic ordering\ instantiation word :: (len) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by (standard; transfer) auto end interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ by (standard; transfer) simp lemma word_of_nat_less_eq_iff: \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_eq_iff: \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_less_iff: \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_iff: \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule lemma word_le_def [code]: "a \ b \ uint a \ uint b" by transfer rule lemma word_less_def [code]: "a < b \ uint a < uint b" by transfer rule lemma word_greater_zero_iff: \a > 0 \ a \ 0\ for a :: \'a::len word\ by transfer (simp add: less_le) lemma of_nat_word_less_eq_iff: \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_nat_word_less_iff: \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_int_word_less_eq_iff: \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma of_int_word_less_iff: \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule subsection \Enumeration\ lemma inj_on_word_of_nat: \inj_on (word_of_nat :: nat \ 'a::len word) {0..<2 ^ LENGTH('a)}\ by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self) lemma UNIV_word_eq_word_of_nat: \(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\ (is \_ = ?A\) proof show \word_of_nat ` {0..<2 ^ LENGTH('a)} \ UNIV\ by simp show \UNIV \ ?A\ proof fix w :: \'a word\ show \w \ (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\ by (rule image_eqI [of _ _ \unat w\]; transfer) simp_all qed qed instantiation word :: (len) enum begin definition enum_word :: \'a word list\ where \enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\ definition enum_all_word :: \('a word \ bool) \ bool\ where \enum_all_word = Ball UNIV\ definition enum_ex_word :: \('a word \ bool) \ bool\ where \enum_ex_word = Bex UNIV\ lemma [code]: \Enum.enum_all P \ Ball UNIV P\ \Enum.enum_ex P \ Bex UNIV P\ for P :: \'a word \ bool\ by (simp_all add: enum_all_word_def enum_ex_word_def) instance by standard (simp_all add: UNIV_word_eq_word_of_nat inj_on_word_of_nat enum_word_def enum_all_word_def enum_ex_word_def distinct_map) end subsection \Bit-wise operations\ instantiation word :: (len) semiring_modulo begin lift_definition divide_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ by simp lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ by simp instance proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed end instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) qed lemma word_bit_induct [case_names zero even odd]: \P a\ if word_zero: \P 0\ and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ for P and a :: \'a::len word\ proof - define m :: nat where \m = LENGTH('a) - Suc 0\ then have l: \LENGTH('a) = Suc m\ by simp define n :: nat where \n = unat a\ then have \n < 2 ^ LENGTH('a)\ by transfer (simp add: take_bit_eq_mod) then have \n < 2 * 2 ^ m\ by (simp add: l) then have \P (of_nat n)\ proof (induction n rule: nat_bit_induct) case zero show ?case by simp (rule word_zero) next case (even n) then have \n < 2 ^ m\ by simp with even.IH have \P (of_nat n)\ by simp moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ by (auto simp add: word_greater_zero_iff l) moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (2 * of_nat n)\ by (rule word_even) then show ?case by simp next case (odd n) then have \Suc n \ 2 ^ m\ by simp with odd.IH have \P (of_nat n)\ by simp moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (1 + 2 * of_nat n)\ by (rule word_odd) then show ?case by simp qed moreover have \of_nat (nat (uint a)) = a\ by transfer simp ultimately show ?thesis by (simp add: n_def) qed lemma bit_word_half_eq: \(of_bool b + a * 2) div 2 = a\ if \a < 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof (cases \2 \ LENGTH('a::len)\) case False have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int by auto with False that show ?thesis by transfer (simp add: eq_iff) next case True obtain n where length: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all show ?thesis proof (cases b) case False moreover have \a * 2 div 2 = a\ using that proof transfer fix k :: int from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ by simp moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ with \LENGTH('a) = Suc n\ have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ by (simp add: take_bit_eq_mod divmod_digit_0) 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 have \
: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" by (metis le_less take_bit_eq_mod take_bit_nonnegative) have less_power: "\n i p. (i::int) mod numeral p ^ n < numeral p ^ n" by simp show \a mod b div b = 0\ for a b :: \'a word\ apply transfer apply (simp add: take_bit_eq_mod mod_eq_0_iff_dvd dvd_def) by (metis (no_types, hide_lams) "\
" Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign le_less_trans mult_eq_0_iff take_bit_eq_mod take_bit_nonnegative zdiv_eq_0_iff zmod_le_nonneg_dividend) show \(1 + a) div 2 = a div 2\ if \even a\ for a :: \'a word\ using that by transfer - (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) + (auto dest: le_Suc_ex simp add: take_bit_Suc elim!: evenE) show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat by transfer (simp, simp add: exp_div_exp_eq) show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" for a :: "'a word" and m n :: nat apply transfer apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) apply (simp add: drop_bit_take_bit) done show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" for a :: "'a word" and m n :: nat by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ if \m \ n\ for a :: "'a word" and m n :: nat using that apply transfer apply (auto simp flip: take_bit_eq_mod) apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) done show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ for a :: "'a word" and m n :: nat by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ for m n :: nat by transfer (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 lemma finite_bit_word [simp]: \finite {n. bit w n}\ for w :: \'a::len word\ proof - have \{n. bit w n} \ {0..LENGTH('a)}\ by (auto dest: bit_imp_le_length) moreover have \finite {0..LENGTH('a)}\ by simp ultimately show ?thesis by (rule finite_subset) qed 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 [code]: \push_bit n w = w * 2 ^ n\ for w :: \'a::len word\ by (fact push_bit_eq_mult) lemma [code]: \Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\ by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) lemma [code]: \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: not_le not_less ac_simps min_absorb2) 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_simps]: \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 unique_euclidean_semiring_with_bit_shifts begin lemma unsigned_drop_bit_eq: \unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\ for w :: \'b::len word\ by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Parity.bit_drop_bit_eq dest: bit_imp_le_length) 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 [simp]: \0 \ unsigned w\ for w :: \'b::len word\ by (transfer fixing: less_eq) simp lemma unsigned_less [simp]: \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ by (transfer fixing: less) simp end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end subsubsection \Generic signed conversion\ context ring_bit_operations begin lemma bit_signed_iff [bit_simps]: \bit (signed w) n \ 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 unat_drop_bit_eq: \unat (drop_bit n w) = drop_bit n (unat w)\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq) lemma uint_mod_distrib: \uint (v mod w) = uint v mod uint w\ proof - have \int (unat (v mod w)) = int (unat v mod unat w)\ by (simp add: unat_mod_distrib) then show ?thesis by (simp add: of_nat_mod) qed context semiring_bit_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 (fact unsigned_greater_eq) lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len word" by (fact unsigned_less) lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len word" by transfer (simp add: take_bit_eq_mod) lemma word_uint_eqI: "uint a = uint b \ a = b" by (fact unsigned_word_eqI) lemma word_uint_eq_iff: "a = b \ uint a = uint b" by (fact word_eq_iff_unsigned) lemma uint_word_of_int_eq: \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer rule lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" by (simp add: uint_word_of_int_eq take_bit_eq_mod) lemma word_of_int_uint: "word_of_int (uint w) = w" by transfer simp lemma word_div_def [code]: "a div b = word_of_int (uint a div uint b)" by transfer rule lemma word_mod_def [code]: "a mod b = word_of_int (uint a mod uint b)" by transfer rule lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" proof fix x :: "'a word" assume "\x. PROP P (word_of_int x)" then have "PROP P (word_of_int (uint x))" . then show "PROP P x" by (simp only: word_of_int_uint) qed lemma sint_uint: \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ for w :: \'a::len word\ by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) lemma unat_eq_nat_uint: \unat w = nat (uint w)\ by simp lemma ucast_eq: \ucast w = word_of_int (uint w)\ by transfer simp lemma scast_eq: \scast w = word_of_int (sint w)\ by transfer simp lemma uint_0_eq: \uint 0 = 0\ by (fact unsigned_0) lemma uint_1_eq: \uint 1 = 1\ by (fact unsigned_1) lemma word_m1_wi: "- 1 = word_of_int (- 1)" by simp lemma uint_0_iff: "uint x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0_iff: "unat x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0: "unat 0 = 0" by (fact unsigned_0) lemma unat_gt_0: "0 < unat x \ x \ 0" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0: "ucast 0 = 0" by (fact unsigned_0) lemma sint_0: "sint 0 = 0" by (fact signed_0) lemma scast_0: "scast 0 = 0" by (fact signed_0) lemma sint_n1: "sint (- 1) = - 1" by (fact signed_minus_1) lemma scast_n1: "scast (- 1) = - 1" by (fact signed_minus_1) lemma uint_1: "uint (1::'a::len word) = 1" by (fact uint_1_eq) lemma unat_1: "unat (1::'a::len word) = 1" by (fact unsigned_1) lemma ucast_1: "ucast (1::'a::len word) = 1" by (fact unsigned_1) instantiation word :: (len) size begin lift_definition size_word :: \'a word \ nat\ is \\_. LENGTH('a)\ .. instance .. end lemma word_size [code]: \size w = LENGTH('a)\ for w :: \'a::len word\ by (fact size_word.rep_eq) lemma word_size_gt_0 [iff]: "0 < size w" for w :: "'a::len word" by (simp add: word_size) lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemma lens_not_0 [iff]: \size w \ 0\ for w :: \'a::len word\ by auto lift_definition source_size :: \('a::len word \ 'b) \ nat\ is \\_. LENGTH('a)\ . lift_definition target_size :: \('a \ 'b::len word) \ nat\ is \\_. LENGTH('b)\ .. lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lemma is_up_eq: \is_up f \ source_size f \ target_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) lemma is_down_eq: \is_down f \ target_size f \ source_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ is \\f. f \ take_bit LENGTH('a)\ by simp lemma word_int_case_eq_uint [code]: \word_int_case f w = f (uint w)\ by transfer simp translations "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" subsection \Arithmetic operations\ text \Legacy theorems:\ lemma word_add_def [code]: "a + b = word_of_int (uint a + uint b)" by transfer (simp add: take_bit_add) lemma word_sub_wi [code]: "a - b = word_of_int (uint a - uint b)" by transfer (simp add: take_bit_diff) lemma word_mult_def [code]: "a * b = word_of_int (uint a * uint b)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_minus_def [code]: "- a = word_of_int (- uint a)" by transfer (simp add: take_bit_minus) lemma word_0_wi: "0 = word_of_int 0" by transfer simp lemma word_1_wi: "1 = word_of_int 1" by transfer simp lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lemma word_succ_alt [code]: "word_succ a = word_of_int (uint a + 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_pred_alt [code]: "word_pred a = word_of_int (uint a - 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemmas word_arith_wis = word_add_def word_sub_wi word_mult_def word_minus_def word_succ_alt word_pred_alt word_0_wi word_1_wi lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" by (transfer, simp)+ lemmas wi_hom_syms = wi_homs [symmetric] lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] lemma double_eq_zero_iff: \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof - define n where \n = LENGTH('a) - Suc 0\ then have *: \LENGTH('a) = Suc n\ by simp have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ using that by transfer (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ by transfer simp then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ by (simp add: *) ultimately show ?thesis by auto qed subsection \Ordering\ lift_definition word_sle :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k \ signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition word_sless :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) notation word_sle ("'(\s')") and word_sle ("(_/ \s _)" [51, 51] 50) and word_sless ("'(a <=s b \ sint a \ sint b\ by transfer simp lemma [code]: \a sint a < sint b\ by transfer simp lemma signed_ordering: \ordering word_sle word_sless\ apply (standard; transfer) using signed_take_bit_decr_length_iff by force+ lemma signed_linorder: \class.linorder word_sle word_sless\ by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) interpretation signed: linorder word_sle word_sless by (fact signed_linorder) lemma word_sless_eq: \x x <=s y \ x \ y\ by (fact signed.less_le) lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len word" by (fact word_coorder.extremum) lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 ) lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len word" by (fact word_order.extremum) lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y \ 0 \ y" for y :: "'a::len word" by (simp add: less_le) lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "a sint a < sint b" by transfer simp lemma word_le_nat_alt: "a \ b \ unat a \ unat b" by transfer (simp add: nat_le_eq_zle) lemma word_less_nat_alt: "a < b \ unat a < unat b" by transfer (auto simp add: less_le [of 0]) lemmas unat_mono = word_less_nat_alt [THEN iffD1] instance word :: (len) wellorder proof fix P :: "'a word \ bool" and a assume *: "(\b. (\a. a < b \ P a) \ P b)" have "wf (measure unat)" .. moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" by (auto simp add: word_less_nat_alt) ultimately have "wf {(a, b :: ('a::len) word). a < b}" by (rule wf_subset) then show "P a" using * by induction blast qed lemma wi_less: "(word_of_int n < (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) lemma wi_le: "(word_of_int n \ (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) subsection \Bit-wise operations\ 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_simps]: \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 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_simps]: \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 lemma bit_word_iff_drop_bit_and [code]: \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) lemma word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" by (transfer, simp add: take_bit_not_take_bit)+ 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_simps]: \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_simps]: \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 \More shift operations\ lift_definition signed_drop_bit :: \nat \ 'a word \ 'a::len word\ is \\n. drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)\ using signed_take_bit_decr_length_iff by (simp add: take_bit_drop_bit) force lemma bit_signed_drop_bit_iff [bit_simps]: \bit (signed_drop_bit m w) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else m + n)\ for w :: \'a::len word\ apply transfer apply (auto simp add: bit_drop_bit_eq bit_signed_take_bit_iff not_le min_def) apply (metis add.commute le_antisym less_diff_conv less_eq_decr_length_iff) apply (metis le_antisym less_eq_decr_length_iff) done lemma [code]: \Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ for w :: \'a::len word\ by transfer simp lemma signed_drop_bit_signed_drop_bit [simp]: \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ for w :: \'a::len word\ proof (cases \LENGTH('a)\) case 0 then show ?thesis using len_not_eq_0 by blast next case (Suc n) then show ?thesis by (force simp add: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps intro!: bit_word_eqI) qed lemma signed_drop_bit_0 [simp]: \signed_drop_bit 0 w = w\ by transfer (simp add: take_bit_signed_take_bit) lemma sint_signed_drop_bit_eq: \sint (signed_drop_bit n w) = drop_bit n (sint w)\ proof (cases \LENGTH('a) = 0 \ n=0\) case False then show ?thesis apply simp apply (rule bit_eqI) by (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) qed auto lift_definition sshiftr1 :: \'a::len word \ 'a word\ is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)\ 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_signed_drop_bit_Suc_0: \sshiftr1 = signed_drop_bit (Suc 0)\ by (rule ext) (transfer, simp add: drop_bit_Suc) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by transfer simp subsection \Rotation\ lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (n mod LENGTH('a)) k)\ subgoal for n k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) done lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (n mod LENGTH('a)) (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ subgoal for n k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) done lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) (take_bit (nat (r mod int LENGTH('a))) k)\ subgoal for r k l by (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) done lemma word_rotl_eq_word_rotr [code]: \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all lemma word_roti_eq_word_rotr_word_rotl [code]: \word_roti i w = (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ proof (cases \i \ 0\) case True moreover define n where \n = nat i\ ultimately have \i = int n\ by simp moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: nat_mod_distrib) ultimately show ?thesis by simp next case False moreover define n where \n = nat (- i)\ ultimately have \i = - int n\ \n > 0\ by simp_all moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) ultimately show ?thesis by simp qed lemma bit_word_rotr_iff [bit_simps]: \bit (word_rotr m w) n \ n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ for w :: \'a::len word\ proof transfer fix k :: int and m n :: nat define q where \q = m mod LENGTH('a)\ have \q < LENGTH('a)\ by (simp add: q_def) then have \q \ LENGTH('a)\ by simp have \m mod LENGTH('a) = q\ by (simp add: q_def) moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ using \q < LENGTH('a)\ by (cases \q + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff le_mod_geq ac_simps) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (m mod LENGTH('a)) k)) n \ n < LENGTH('a) \ (n + m) mod LENGTH('a) < LENGTH('a) \ bit k ((n + m) mod LENGTH('a))\ by simp qed lemma bit_word_rotl_iff [bit_simps]: \bit (word_rotl m w) n \ n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ for w :: \'a::len word\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) lemma bit_word_roti_iff [bit_simps]: \bit (word_roti k w) n \ n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ for w :: \'a::len word\ proof transfer fix k l :: int and n :: nat define m where \m = nat (k mod int LENGTH('a))\ have \m < LENGTH('a)\ by (simp add: nat_less_iff m_def) then have \m \ LENGTH('a)\ by simp have \k mod int LENGTH('a) = int m\ by (simp add: nat_less_iff m_def) moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ using \m < LENGTH('a)\ by (cases \m + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff nat_less_iff not_le not_less ac_simps le_diff_conv le_mod_geq) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) (take_bit (nat (k mod int LENGTH('a))) l)) n \ n < LENGTH('a) \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) \ bit l (nat ((int n + k) mod int LENGTH('a)))\ by simp qed lemma uint_word_rotr_eq: \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (uint w)) (uint (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ apply transfer by (simp add: min.absorb2 take_bit_concat_bit_eq) lemma [code]: \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (Word.the_int w)) (Word.the_int (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ using uint_word_rotr_eq [of n w] by simp subsection \Split and cat operations\ lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) lemma word_cat_eq: \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: concat_bit_eq ac_simps) lemma word_cat_eq' [code]: \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ for a :: \'a::len word\ and b :: \'b::len word\ by transfer (simp add: concat_bit_take_bit_eq) lemma bit_word_cat_iff [bit_simps]: \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) definition word_split :: \'a::len word \ 'b::len word \ 'c::len word\ where \word_split w = (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\ definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ abbreviation (input) max_word :: \'a::len word\ \ \Largest representable machine integer.\ where "max_word \ - 1" subsection \More on conversions\ lemma int_word_sint: \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift) lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by simp lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)" for w :: "'a::len word" by transfer (simp add: take_bit_signed_take_bit) lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" by transfer (simp add: min_def) lemma wi_bintr: "LENGTH('a::len) \ n \ word_of_int (take_bit n w) = (word_of_int w :: 'a word)" by transfer simp lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" by (induct b, simp_all only: numeral.simps word_of_int_homs) declare word_numeral_alt [symmetric, code_abbrev] lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" by (simp only: word_numeral_alt wi_hom_neg) declare word_neg_numeral_alt [symmetric, code_abbrev] lemma uint_bintrunc [simp]: "uint (numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (numeral bin)" by transfer rule lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" by transfer rule lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" by transfer simp lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" by transfer simp lemma unat_bintrunc [simp]: "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" by transfer simp lemma unat_bintrunc_neg [simp]: "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" by transfer simp lemma size_0_eq: "size w = 0 \ v = w" for v w :: "'a::len word" by transfer simp lemma uint_ge_0 [iff]: "0 \ uint x" by (fact unsigned_greater_eq) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len word" by (fact unsigned_less) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using sint_greater_eq [of x] by simp lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using sint_less [of x] by simp lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len word" by (simp only: diff_less_0_iff_less uint_lt2p) lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" for x :: "'a::len word" by (simp only: not_le uint_m2p_neg) lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len word" using uint_bounded [of w] by (rule less_le_trans) simp lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) lemma uint_nat: "uint w = int (unat w)" by transfer simp lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma sint_numeral: "sint (numeral b :: 'a::len word) = (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" by (metis int_word_sint word_numeral_alt) lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" by (fact of_int_0) lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" by (fact of_int_1) lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" by (simp add: wi_hom_syms) lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" by (fact of_int_numeral) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" by (fact of_int_neg_numeral) lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" by transfer (simp add: take_bit_eq_mod) lemma word_int_split: "P (word_int_case f x) = (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" by transfer (auto simp add: take_bit_eq_mod) lemma word_int_split_asm: "P (word_int_case f x) = (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" by transfer (auto simp add: take_bit_eq_mod) lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" by transfer simp lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" by (simp add: word_size sint_greater_eq sint_less) lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" unfolding word_size by (rule less_le_trans [OF sint_lt]) lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" for w :: "'a::len word" unfolding word_size by (rule order_trans [OF _ sint_ge]) subsection \Testing bits\ lemma bin_nth_uint_imp: "bit (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" by transfer (simp add: bit_take_bit_iff) lemma bin_nth_sint: "LENGTH('a) \ n \ bit (sint w) n = bit (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) lemma num_of_bintr': "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \take_bit LENGTH('a) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by simp qed lemma num_of_sbintr': "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by (simp add: take_bit_signed_take_bit) qed lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (take_bit (LENGTH('a::len)) (numeral x))" by transfer simp lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" by transfer (simp add: take_bit_signed_take_bit) text \ \cast\ -- note, no arg for new length, as it's determined by type of result, thus in \cast w = w\, the type means cast to length of \w\! \ lemma bit_ucast_iff: \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ 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 add: take_bit_signed_take_bit) lemma ucast_mask_eq: \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff 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 word_log_defs = word_and_def word_or_def word_xor_def word_not_def lemma bit_last_iff: \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) for w :: \'a::len word\ proof - have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ by (simp add: bit_uint_iff) also have \\ \ ?Q\ by (simp add: sint_uint) finally show ?thesis . qed lemma drop_bit_eq_zero_iff_not_bit_last: \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ for w :: "'a::len word" proof (cases \LENGTH('a)\) case (Suc n) then show ?thesis apply transfer apply (simp add: take_bit_drop_bit) by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit min.absorb2 odd_iff_mod_2_eq_one) qed auto subsection \Word Arithmetic\ 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)" by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) lemma uint_word_arith_bintrs: fixes a b :: "'a::len word" shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" by (simp_all add: uint_word_ariths take_bit_eq_mod) lemma sint_word_ariths: fixes a b :: "'a::len word" shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" subgoal by transfer (simp add: signed_take_bit_add) subgoal by transfer (simp add: signed_take_bit_diff) subgoal by transfer (simp add: signed_take_bit_mult) subgoal by transfer (simp add: signed_take_bit_minus) apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ) apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred) apply (simp_all add: sint_uint) done lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" unfolding word_pred_m1 by simp lemma succ_pred_no [simp]: "word_succ (numeral w) = numeral w + 1" "word_pred (numeral w) = numeral w - 1" "word_succ (- numeral w) = - numeral w + 1" "word_pred (- numeral w) = - numeral w - 1" by (simp_all add: word_succ_p1 word_pred_m1) lemma word_sp_01 [simp]: "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" by (simp_all add: word_succ_p1 word_pred_m1) \ \alternative approach to lifting arithmetic equalities\ lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp subsection \Order on fixed-length words\ lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp lemma udvd_iff_dvd: \x udvd y \ unat x dvd unat y\ by transfer (simp add: nat_dvd_iff) lemma udvd_iff_dvd_int: \v udvd w \ uint v dvd uint w\ by transfer rule lemma udvdI [intro]: \v udvd w\ if \unat w = unat v * unat u\ proof - from that have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvdE [elim]: fixes v w :: \'a::len word\ assumes \v udvd w\ obtains u :: \'a word\ where \unat w = unat v * unat u\ proof (cases \v = 0\) case True moreover from True \v udvd w\ have \w = 0\ by transfer simp ultimately show thesis using that by simp next case False then have \unat v > 0\ by (simp add: unat_gt_0) from \v udvd w\ have \unat v dvd unat w\ by (simp add: udvd_iff_dvd) then obtain n where \unat w = unat v * n\ .. moreover have \n < 2 ^ LENGTH('a)\ proof (rule ccontr) assume \\ n < 2 ^ LENGTH('a)\ then have \n \ 2 ^ LENGTH('a)\ by (simp add: not_le) then have \unat v * n \ 2 ^ LENGTH('a)\ using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] by simp with \unat w = unat v * n\ have \unat w \ 2 ^ LENGTH('a)\ by simp with unsigned_less [of w, where ?'a = nat] show False by linarith qed ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ by (auto simp add: take_bit_nat_eq_self_iff intro: sym) with that show thesis . qed lemma udvd_imp_mod_eq_0: \w mod v = 0\ if \v udvd w\ using that by transfer simp lemma mod_eq_0_imp_udvd [intro?]: \v udvd w\ if \w mod v = 0\ proof - from that have \unat (w mod v) = unat 0\ by simp then have \unat w mod unat v = 0\ by (simp add: unat_mod_distrib) then have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvd_imp_dvd: \v dvd w\ if \v udvd w\ for v w :: \'a::len word\ proof - from that obtain u :: \'a word\ where \unat w = unat v * unat u\ .. then have \(word_of_nat (unat w) :: 'a word) = word_of_nat (unat v * unat u)\ by simp then have \w = v * u\ by simp then show \v dvd w\ .. qed lemma exp_dvd_iff_exp_udvd: \2 ^ n dvd w \ 2 ^ n udvd w\ for v w :: \'a::len word\ proof assume \2 ^ n udvd w\ then show \2 ^ n dvd w\ by (rule udvd_imp_dvd) next assume \2 ^ n dvd w\ then obtain u :: \'a word\ where \w = 2 ^ n * u\ .. then have \w = push_bit n u\ by (simp add: push_bit_eq_mult) then show \2 ^ n udvd w\ by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod) qed lemma udvd_nat_alt: \a udvd b \ (\n. unat b = n * unat a)\ by (auto simp add: udvd_iff_dvd) lemma udvd_unfold_int: \a udvd b \ (\n\0. uint b = n * uint a)\ unfolding udvd_iff_dvd_int by (metis dvd_div_mult_self dvd_triv_right uint_div_distrib uint_ge_0) lemma unat_minus_one: \unat (w - 1) = unat w - 1\ if \w \ 0\ proof - have "0 \ uint w" by (fact uint_nonnegative) moreover from that have "0 \ uint w" by (simp add: uint_0_iff) ultimately have "1 \ uint w" by arith from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" by arith with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" by (auto intro: mod_pos_pos_trivial) with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" by (auto simp del: nat_uint_eq) then show ?thesis by (simp only: unat_eq_nat_uint word_arith_wis mod_diff_right_eq) (metis of_int_1 uint_word_of_int unsigned_1) qed lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" for x :: "'a::len word" and y :: "'b::len word" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection \Conditions for the addition (etc) of two words to overflow\ lemma uint_add_lem: "(uint x + uint y < 2 ^ LENGTH('a)) = (uint (x + y) = uint x + uint y)" for x y :: "'a::len word" by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) lemma uint_mult_lem: "(uint x * uint y < 2 ^ LENGTH('a)) = (uint (x * y) = uint x * uint y)" for x y :: "'a::len word" by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" unfolding uint_word_ariths by (simp flip: take_bit_eq_mod add: take_bit_int_greater_eq_self_iff) lemma int_mod_ge: \a \ a mod n\ if \a < n\ \0 < n\ for a n :: int proof (cases \a < 0\) case True with \0 < n\ show ?thesis by (metis less_trans not_less pos_mod_conj) next case False with \a < n\ show ?thesis by simp qed lemma mod_add_if_z: "\x < z; y < z; 0 \ y; 0 \ x; 0 \ z\ \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" for x y z :: int apply (simp add: not_less) by (metis (no_types) add_strict_mono diff_ge_0_iff_ge diff_less_eq minus_mod_self2 mod_pos_pos_trivial) lemma uint_plus_if': "uint (a + b) = (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b else uint a + uint b - 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) lemma mod_sub_if_z: "\x < z; y < z; 0 \ y; 0 \ x; 0 \ z\ \ (x - y) mod z = (if y \ x then x - y else x - y + z)" for x y z :: int using mod_pos_pos_trivial [of "x - y + z" z] by (auto simp add: not_le) lemma uint_sub_if': "uint (a - b) = (if uint b \ uint a then uint a - uint b else uint a - uint b + 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) 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" by transfer (simp add: take_bit_int_eq_self) lemma uint_split: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" for x :: "'a::len word" by transfer (auto simp add: take_bit_eq_mod) lemma uint_split_asm: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" for x :: "'a::len word" by auto (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_eq_iff 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 \ val uint_arith_simpset = @{context} |> fold Simplifier.add_simp @{thms uint_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of; fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset uint_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms uint_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n THEN assume_tac ctxt n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) \ method_setup uint_arith = \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ "solving word arithmetic via integers and arith" subsection \More on overflows and monotonicity\ lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" for x y :: "'a::len word" 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" unfolding word_less_alt word_le_def by (metis (full_types) order_trans udvd_incr_lem uint_add_le) lemma udvd_decr': assumes "p < q" "uint p = ua + n * uint K" "uint q = ua + n' * uint K" shows "uint q = ua + n' * uint K \ p \ q - K" proof - have "\w wa. uint (w::'a word) \ uint wa + uint (w - wa)" by (metis (no_types) add_diff_cancel_left' diff_add_cancel uint_add_le) moreover have "uint K + uint p \ uint q" using assms by (metis (no_types) add_diff_cancel_left' diff_add_cancel udvd_incr_lem word_less_def) ultimately show ?thesis by (meson add_le_cancel_left order_trans word_less_eq_iff_unsigned) qed lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" unfolding udvd_unfold_int by (meson udvd_decr0) lemma udvd_incr2_K: "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ 0 < K \ p \ p + K \ p + K \ a + s" unfolding udvd_unfold_int apply (simp add: uint_arith_simps split: if_split_asm) apply (metis (no_types, hide_lams) le_add_diff_inverse le_less_trans udvd_incr_lem) using uint_lt2p [of s] by simp subsection \Arithmetic type class instantiations\ lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN antisym_conv1] lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" by simp text \ note that \iszero_def\ is only for class \comm_semiring_1_cancel\, which requires word length \\ 1\, ie \'a::len word\ \ lemma iszero_word_no [simp]: "iszero (numeral bin :: 'a::len word) = iszero (take_bit LENGTH('a) (numeral bin :: int))" by (metis iszero_def uint_0_iff uint_bintrunc) text \Use \iszero\ to simplify equalities between word numerals.\ lemmas word_eq_numeral_iff_iszero [simp] = eq_numeral_iff_iszero [where 'a="'a::len word"] subsection \Word and nat\ lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" by (metis of_nat_unat ucast_id unsigned_less) lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" for w :: "'a::len word" using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] by (auto simp flip: take_bit_eq_mod) lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" by (simp add: of_nat_eq) lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" by (auto simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" by (simp add: wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by transfer (simp add: ac_simps) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by simp lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by simp lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by simp lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemma unat_cong: "x = y \ unat x = unat y" by (fact arg_cong) lemma unat_of_nat: \unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\ by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq) lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat]] lemmas word_sub_less_iff = word_sub_le_iff [unfolded linorder_not_less [symmetric] Not_eq_iff] lemma unat_add_lem: "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" for x y :: "'a::len word" by (metis mod_less unat_word_ariths(1) unsigned_less) lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" by (metis mod_less unat_word_ariths(2) unsigned_less) lemma 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 le_iff_add) by (metis add.commute add_less_cancel_right add_strict_mono mod_less unsigned_less) lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len word" using word_le_plus_either by blast 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)" proof - { assume xy: "\ uint y \ uint x" have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x + 2 ^ LENGTH('a) - uint y)" by simp also have "... = nat (uint x + 2 ^ LENGTH('a)) - nat (uint y)" by (simp add: nat_diff_distrib') also have "... = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" by (metis nat_add_distrib nat_eq_numeral_power_cancel_iff order_less_imp_le unsigned_0 unsigned_greater_eq unsigned_less) finally have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" . } then show ?thesis unfolding word_size by (metis nat_diff_distrib' uint_sub_if' un_ui_le unat_eq_nat_uint unsigned_greater_eq) qed lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma uint_div: \uint (x div y) = uint x div uint y\ by (fact uint_div_distrib) lemma unat_div: \unat (x div y) = unat x div unat y\ by (fact unat_div_distrib) lemma uint_mod: \uint (x mod y) = uint x mod uint y\ by (fact uint_mod_distrib) lemma unat_mod: \unat (x mod y) = unat x mod unat y\ 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 (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 (metis take_bit_nat_eq_self_iff) lemma of_nat_inverse: \word_of_nat r = a \ r < 2 ^ LENGTH('a) \ unat a = r\ for a :: \'a::len word\ by (metis mod_if unat_of_nat) lemma word_unat_eq_iff: \v = w \ unat v = unat w\ for v w :: \'a::len word\ by (fact word_eq_iff_unsigned) lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat_eq_iff unat_sub_if' unat_plus_if' unat_div unat_mod \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ val unat_arith_simpset = @{context} (* TODO: completely explicitly determined simpset *) |> fold Simplifier.del_simp @{thms unsigned_of_nat unsigned_of_int} |> fold Simplifier.add_simp @{thms unat_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset unat_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms unat_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) \ method_setup unat_arith = \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ "solving word arithmetic via natural numbers and arith" lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem] lemma word_div_mult: "\0 < y; unat x * unat y < 2 ^ LENGTH('a)\ \ x * y div y = x" for x y :: "'a::len word" by (simp add: unat_eq_zero unat_mult_lem word_arith_nat_div) lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" by unat_arith (meson le_less_trans less_mult_imp_div_less not_le unsigned_less) lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "\i < k div x; 0 < x\ \ i * x < k" for i k x :: "'a::len word" by (metis div_le_mono div_lt'' not_le unat_div word_div_mult word_less_iff_unsigned) lemma div_le_mult: "\i \ k div x; 0 < x\ \ i * x \ k" for i k x :: "'a::len word" by (metis div_lt' less_mult_imp_div_less not_less unat_arith_simps(2) unat_div unat_mult_lem) lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" unfolding uint_nat by (metis div_lt' int_ops(7) of_nat_unat uint_mult_lem unat_mult_lem) lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" for x y z :: "'a::len word" by (metis add.commute diff_add_cancel no_olen_add) lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x lemma le_unat_uoi: \y \ unat z \ unat (word_of_nat y :: 'a word) = y\ for z :: \'a::len word\ by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans) lemmas thd = times_div_less_eq_dividend lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" for n b :: "'a::len word" by (fact div_mult_mod_eq) lemma word_div_mult_le: "a div b * b \ a" for a b :: "'a::len word" by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) lemma word_mod_less_divisor: "0 < n \ m mod n < n" for m n :: "'a::len word" by (simp add: unat_arith_simps) lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" by (induct n) (simp_all add: wi_hom_mult [symmetric]) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" for n :: "'a::len word" by unat_arith subsection \Cardinality, finiteness of set of words\ lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ unfolding inj_on_def by (metis atLeastLessThan_iff word_of_int_inverse) lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ apply transfer apply (auto simp add: image_iff) apply (metis take_bit_int_eq_self_iff) done lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ by (auto simp add: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split) lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" by (simp add: UNIV_eq card_image inj_on_word_of_int) lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len word" unfolding word_size by (rule card_word) instance word :: (len) finite by standard (simp add: UNIV_eq) subsection \Bitwise Operations on Words\ 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 \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add word_wi_log_defs lemma word_bw_assocs: "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" for x :: "'a::len word" by (fact ac_simps)+ lemma word_bw_comms: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: "'a::len word" by (fact ac_simps)+ lemma word_bw_lcs: "y AND x AND z = x AND y AND z" "y OR x OR z = x OR y OR z" "y XOR x XOR z = x XOR y XOR z" for x :: "'a::len word" by (fact ac_simps)+ lemma word_log_esimps: "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" for x :: "'a::len word" by simp_all lemma word_not_dist: "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" for x :: "'a::len word" by simp_all lemma word_bw_same: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: "'a::len word" by simp_all lemma word_ao_absorbs [simp]: "x AND (y OR x) = x" "x OR y AND x = x" "x AND (x OR y) = x" "y AND x OR x = x" "(y OR x) AND x = x" "x OR x AND y = x" "(x OR y) AND x = x" "x AND y OR x = x" for x :: "'a::len word" by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" by (fact bit.double_compl) lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len word" by (fact bit.conj_disj_distrib2) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len word" by (fact bit.disj_conj_distrib2) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" by (simp add: not_eq_complement) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len word" by transfer (simp add: plus_and_or) lemma leoa: "w = x OR y \ y = w AND y" for x :: "'a::len word" by auto lemma leao: "w' = x' AND y' \ x' = x' OR w'" for x' :: "'a::len word" by auto lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" for w w' :: "'a::len word" by (auto intro: leoa leao) lemma le_word_or2: "x \ x OR y" for x y :: "'a::len word" by (simp add: or_greater_eq uint_or word_le_def) lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] lemma bit_horner_sum_bit_word_iff [bit_simps]: \bit (horner_sum of_bool (2 :: 'a::len word) bs) n \ n < min LENGTH('a) (length bs) \ bs ! n\ by transfer (simp add: bit_horner_sum_bit_iff) definition word_reverse :: \'a::len word \ 'a word\ where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. lemma bit_word_reverse_iff [bit_simps]: \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ for w :: \'a::len word\ by (cases \n < LENGTH('a)\) (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" by (cases \n < LENGTH('a)\; transfer; force) lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) 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 text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), where \f\ (ie \_ div 2\) takes normal arguments to normal results, thus we get (2) from (1) \ lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" using drop_bit_eq_div [of 1 \uint w\, symmetric] by transfer (simp add: drop_bit_take_bit min_def) lemma bit_sshiftr1_iff [bit_simps]: \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 by (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def le_Suc_eq simp flip: bit_Suc) 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" using sint_signed_drop_bit_eq [of 1 w] by (simp add: drop_bit_Suc sshiftr1_eq_signed_drop_bit_Suc_0) lemma bit_bshiftr1_iff [bit_simps]: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ apply transfer 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))" by (intro bit_word_eqI) (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) \ \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 shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (take_bit LENGTH('a) (numeral w) div 2)" by transfer simp lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" by transfer simp 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 False_map2_or: "\set xs \ {False}; length ys = length xs\ \ map2 (\) xs ys = ys" by (induction xs arbitrary: ys) (auto simp: length_Suc_conv) lemma align_lem_or: assumes "length xs = n + m" "length ys = n + m" and "drop m xs = replicate n False" "take m ys = replicate m False" shows "map2 (\) xs ys = take m xs @ drop m ys" using assms proof (induction xs arbitrary: ys m) case (Cons a xs) then show ?case by (cases m) (auto simp: length_Suc_conv False_map2_or) qed auto lemma False_map2_and: "\set xs \ {False}; length ys = length xs\ \ map2 (\) xs ys = xs" by (induction xs arbitrary: ys) (auto simp: length_Suc_conv) lemma align_lem_and: assumes "length xs = n + m" "length ys = n + m" and "drop m xs = replicate n False" "take m ys = replicate m False" shows "map2 (\) xs ys = replicate (n + m) False" using assms proof (induction xs arbitrary: ys m) case (Cons a xs) then show ?case by (cases m) (auto simp: length_Suc_conv set_replicate_conv_if False_map2_and) qed auto subsubsection \Mask\ lemma minus_1_eq_mask: \- 1 = (mask LENGTH('a) :: 'a::len word)\ by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff 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_simps]: \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 mask_bin: "mask n = word_of_int (take_bit n (- 1))" by transfer (simp add: take_bit_minus_one_eq_mask) lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" by transfer (simp add: ac_simps take_bit_eq_mask) lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" by (auto simp add: and_mask_bintr min_def not_le wi_bintr) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" by (auto simp add: and_mask_wi min_def wi_bintr) lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" by (simp only: and_mask_bintr take_bit_eq_mod) lemma uint_mask_eq: \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" by (metis take_bit_eq_mask take_bit_int_less_exp unsigned_take_bit_eq) lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (auto simp flip: take_bit_eq_mask) apply (metis take_bit_int_eq_self_iff uint_take_bit_eq) apply (simp add: take_bit_int_eq_self unsigned_take_bit_eq word_uint_eqI) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" for w :: "'a::len word" by transfer simp lemma less_mask_eq: fixes x :: "'a::len word" assumes "x < 2 ^ n" shows "x AND mask n = x" by (metis (no_types) assms lt2p_lem mask_eq_iff not_less word_2p_lem word_size) lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" for x :: \'a::len word\ unfolding word_size by (erule and_mask_less') lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" for c x :: "'a::len word" by (auto simp: word_mod_def uint_2p and_mask_mod_2p) lemma mask_eqs: "(a AND mask n) + b AND mask n = a + b AND mask n" "a + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - b AND mask n = a - b AND mask n" "a - (b AND mask n) AND mask n = a - b AND mask n" "a * (b AND mask n) AND mask n = a * b AND mask n" "(b AND mask n) * a AND mask n = b * a AND mask n" "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" "- (a AND mask n) AND mask n = - a AND mask n" "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] unfolding take_bit_eq_mask [symmetric] by (transfer; simp add: take_bit_eq_mod mod_simps)+ lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" for x :: \'a::len word\ using word_of_int_Ex [where x=x] unfolding take_bit_eq_mask [symmetric] by (transfer; simp add: take_bit_eq_mod mod_simps)+ lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by transfer (simp 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_simps]: \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ for w :: \'a::len word\ by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff not_less not_le ac_simps dest: bit_imp_le_length) definition slice :: \nat \ 'a::len word \ 'b::len word\ where \slice n = slice1 (LENGTH('a) - n)\ lemma bit_slice_iff [bit_simps]: \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ for w :: \'a::len word\ by (simp add: slice_def word_size bit_slice1_iff) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp lemma slice_0 [simp] : "slice n 0 = 0" unfolding slice_def by auto lemma ucast_slice1: "ucast w = slice1 (size w) w" unfolding slice1_def by (simp add: size_word.rep_eq) lemma ucast_slice: "ucast w = slice 0 w" by (simp add: slice_def slice1_def) lemma slice_id: "slice 0 t = t" by (simp only: ucast_slice [symmetric] ucast_id) lemma rev_slice1: \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ if \n + k = LENGTH('a) + LENGTH('b)\ proof (rule bit_word_eqI) fix m assume *: \m < LENGTH('a)\ from that have **: \LENGTH('b) = n + k - LENGTH('a)\ by simp show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ unfolding bit_slice1_iff bit_word_reverse_iff using * ** by (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) auto qed lemma rev_slice: "n + k + LENGTH('a::len) = LENGTH('b::len) \ slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" unfolding slice_def word_size by (simp add: rev_slice1) subsubsection \Revcast\ definition revcast :: \'a::len word \ 'b::len word\ where \revcast = slice1 LENGTH('b)\ lemma bit_revcast_iff [bit_simps]: \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ for w :: \'a::len word\ by (simp add: revcast_def bit_slice1_iff) lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" by (simp add: revcast_def word_size) lemma revcast_rev_ucast [OF refl refl refl]: "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ rc = word_reverse uc" by (metis rev_slice1 revcast_slice1 ucast_slice1 word_size) lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" using revcast_rev_ucast [of "word_reverse w"] by simp lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" by (fact revcast_rev_ucast [THEN word_rev_gal']) lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" by (fact revcast_ucast [THEN word_rev_gal']) text "linking revcast and cast via shift" lemmas wsst_TYs = source_size target_size word_size lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_eq \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" by transfer (simp add: take_bit_concat_bit_eq) lemma word_cat_split_alt: "\size w \ size u + size v; word_split w = (u,v)\ \ word_cat u v = w" unfolding word_split_def by (rule bit_word_eqI) (auto simp add: bit_word_cat_iff not_less word_size bit_ucast_iff bit_drop_bit_eq) lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: assumes "word_split w = (u, v)" shows "u = slice (size v) w \ v = slice 0 w" unfolding word_size proof (intro conjI) have \
: "\n. \ucast (drop_bit LENGTH('b) w) = u; LENGTH('c) < LENGTH('b)\ \ \ bit u n" by (metis bit_take_bit_iff bit_word_of_int_iff diff_is_0_eq' drop_bit_take_bit less_imp_le less_nat_zero_code of_int_uint unsigned_drop_bit_eq) show "u = slice LENGTH('b) w" proof (rule bit_word_eqI) show "bit u n = bit ((slice LENGTH('b) w)::'a word) n" if "n < LENGTH('a)" for n using assms bit_imp_le_length unfolding word_split_def bit_slice_iff by (fastforce simp add: \
ac_simps word_size bit_ucast_iff bit_drop_bit_eq) qed show "v = slice 0 w" by (metis Pair_inject assms ucast_slice word_split_bin') qed lemma slice_cat1 [OF refl]: "\wc = word_cat a b; size a + size b \ size wc\ \ slice (size b) wc = a" by (rule bit_word_eqI) (auto simp add: bit_slice_iff bit_word_cat_iff word_size) lemmas slice_cat2 = trans [OF slice_id word_cat_id] lemma cat_slices: "\a = slice n c; b = slice 0 c; n = size b; size c \ size a + size b\ \ word_cat a b = c" by (rule bit_word_eqI) (auto simp add: bit_slice_iff bit_word_cat_iff word_size) lemma word_split_cat_alt: assumes "w = word_cat u v" and size: "size u + size v \ size w" shows "word_split w = (u,v)" proof - have "ucast ((drop_bit LENGTH('c) (word_cat u v))::'a word) = u" "ucast ((word_cat u v)::'a word) = v" using assms by (auto simp add: word_size bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff intro: bit_eqI) then show ?thesis by (simp add: assms(1) word_split_bin') qed lemma horner_sum_uint_exp_Cons_eq: \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ for ws :: \'a::len word list\ by (simp add: bintr_uint concat_bit_eq push_bit_eq_mult) lemma bit_horner_sum_uint_exp_iff: \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ for ws :: \'a::len word list\ proof (induction ws arbitrary: n) case Nil then show ?case by simp next case (Cons w ws) then show ?case by (cases \n \ LENGTH('a)\) (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) qed subsection \Rotation\ lemma word_rotr_word_rotr_eq: \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) lemma word_rot_lem: "\l + k = d + k mod l; n < l\ \ ((d + n) mod l) = n" for l::nat by (metis (no_types, lifting) add.commute add.right_neutral add_diff_cancel_left' mod_if mod_mult_div_eq mod_mult_self2 mod_self) lemma word_rot_rl [simp]: \word_rotl k (word_rotr k v) = v\ proof (rule bit_word_eqI) show "bit (word_rotl k (word_rotr k v)) n = bit v n" if "n < LENGTH('a)" for n using that by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split) qed lemma word_rot_lr [simp]: \word_rotr k (word_rotl k v) = v\ proof (rule bit_word_eqI) show "bit (word_rotr k (word_rotl k v)) n = bit v n" if "n < LENGTH('a)" for n using that by (auto simp add: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split) qed lemma word_rot_gal: \word_rotr n v = w \ word_rotl n w = v\ by auto lemma word_rot_gal': \w = word_rotr n v \ v = word_rotl n w\ by auto lemma word_rotr_rev: \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('a)\ moreover have \1 + ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = int LENGTH('a)\ apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod int LENGTH('a) = 0\) using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] apply simp_all apply (auto simp add: algebra_simps) apply (metis (mono_tags, hide_lams) Abs_fnat_hom_add mod_Suc mod_mult_self2_is_0 of_nat_Suc of_nat_mod semiring_char_0_class.of_nat_neq_0) apply (metis (no_types, 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" by (rule bit_word_eqI, auto simp add: bit_word_rotl_iff bit_word_rotr_iff bit_and_iff bit_or_iff bit_xor_iff bit_not_iff algebra_simps not_le)+ end 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 (rule that [of \uint x\]) simp_all lemma word_nat_cases [cases type: word]: fixes x :: "'a::len word" obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" by (rule that [of \unat x\]) simp_all lemma max_word_max [intro!]: "n \ max_word" by (fact word_order.extremum) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" by simp lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" by (fact word_exp_length_eq_0) lemma max_word_wrap: "x + 1 = 0 \ x = max_word" by (simp add: eq_neg_iff_add_eq_0) 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 (fact bit.conj_disj_distrib) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len word" by (fact bit.disj_conj_distrib) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len word" by (fact bit.conj_cancel_right) lemma word_or_not [simp]: "x OR NOT x = max_word" by (fact bit.disj_cancel_right) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len word" by (fact bit.xor_def) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" by transfer simp lemma word_less_1 [simp]: "x < 1 \ x = 0" for x :: "'a::len word" by (simp add: word_less_nat_alt unat_0_iff) lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" by (simp add: take_bit_eq_mod word_size uint_word_of_int_eq uint_plus_if') lemma unat_plus_if_size: "unat (x + y) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" for x y :: "'a::len word" by (simp add: size_word.rep_eq unat_arith_simps) lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" by (fact word_coorder.not_eq_extremum) lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" for c :: "'a::len word" by (fact unat_div) lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" by (simp add: size_word.rep_eq uint_sub_if') lemma unat_sub: \unat (a - b) = unat a - unat b\ if \b \ a\ by (meson that unat_sub_if_size word_le_nat_alt) lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" by simp lemma word_of_int_inj: \(word_of_int x :: 'a::len word) = word_of_int y \ x = y\ if \0 \ x \ x < 2 ^ LENGTH('a)\ \0 \ y \ y < 2 ^ LENGTH('a)\ using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) lemma word_le_less_eq: "x \ y \ x = y \ x < y" for x y :: "'z::len word" by (auto simp add: order_class.le_less) lemma mod_plus_cong: fixes b b' :: int assumes 1: "b = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: fixes b b' :: int assumes "b = b'" and "x mod b' = x' mod b'" and "y mod b' = y' mod b'" and "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms [symmetric] by (auto intro: mod_diff_cong) lemma word_induct_less [case_names zero less]: \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ for m :: \'a::len word\ proof - define q where \q = unat m\ with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ by simp then have \P (word_of_nat q :: 'a word)\ proof (induction q) case 0 show ?case by (simp add: zero) next case (Suc q) show ?case proof (cases \1 + word_of_nat q = (0 :: 'a word)\) case True then show ?thesis by (simp add: zero) next case False then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ by (simp add: unatSuc word_less_nat_alt) then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) have \P (word_of_nat q)\ by (simp add: "**" Suc.IH Suc.prems) with * have \P (1 + word_of_nat q)\ by (rule Suc.prems) then show ?thesis by simp qed qed with \q = unat m\ show ?thesis by simp qed lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" by (rule word_induct_less) lemma word_induct2 [case_names zero suc, induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" by (induction rule: word_induct_less; force) subsection \Recursion combinator for words\ definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0 [simp]: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc [simp]: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" by (simp add: unatSuc word_rec_def) lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" by (metis add.commute diff_add_cancel word_rec_Suc) lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: 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_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" proof (induction n arbitrary: z f) case zero then show ?case by (metis diff_0_right word_le_0_iff word_rec_0) next case (suc n z f) show ?case proof (cases "1 + (n - m) = 0") case True then show ?thesis by (simp add: add_diff_eq) next case False then have eq: "1 + n - m = 1 + (n - m)" by simp with False have "m \ n" by (metis "suc.prems" add.commute dual_order.antisym eq_iff_diff_eq_0 inc_le leI) with False "suc.hyps" show ?thesis using suc.IH [of "f 0 z" "f \ (+) 1"] by (simp add: word_rec_in2 eq add.assoc o_def) qed qed lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) auto lemma word_rec_id_eq: "(\m. m < n \ f m = id) \ word_rec z f n = z" by (induction n) (auto simp add: unatSuc unat_arith_simps(2)) lemma word_rec_max: assumes "\m\n. m \ - 1 \ f m = id" shows "word_rec z f (- 1) = word_rec z f n" proof - have \
: "\m. \m < - 1 - n\ \ (f \ (+) n) m = id" using assms by (metis (mono_tags, lifting) add.commute add_diff_cancel_left' comp_apply less_le olen_add_eqv plus_minus_no_overflow word_n1_ge) have "word_rec z f (- 1) = word_rec (word_rec z f (- 1 - (- 1 - n))) (f \ (+) (- 1 - (- 1 - n))) (- 1 - n)" by (meson word_n1_ge word_rec_twice) also have "... = word_rec z f n" by (metis (no_types, lifting) \
diff_add_cancel minus_diff_eq uminus_add_conv_diff word_rec_id_eq) finally show ?thesis . qed subsection \More\ lemma mask_1: "mask 1 = 1" by simp lemma mask_Suc_0: "mask (Suc 0) = 1" by simp lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd n" by simp lemma push_bit_word_beyond [simp]: \push_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by (transfer fixing: n) (simp add: take_bit_push_bit) lemma drop_bit_word_beyond [simp]: \drop_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by (transfer fixing: n) (simp add: drop_bit_take_bit) lemma signed_drop_bit_beyond: \signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\ if \LENGTH('a) \ n\ for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that) subsection \SMT support\ ML_file \Tools/smt_word.ML\ end diff --git a/src/HOL/Library/Z2.thy b/src/HOL/Library/Z2.thy --- a/src/HOL/Library/Z2.thy +++ b/src/HOL/Library/Z2.thy @@ -1,237 +1,237 @@ (* Title: HOL/Library/Z2.thy Author: Brian Huffman *) section \The Field of Integers mod 2\ theory Z2 imports Main "HOL-Library.Bit_Operations" begin text \ Note that in most cases \<^typ>\bool\ is appropriate hen a binary type is needed; the type provided here, for historical reasons named \<^text>\bit\, is only needed if proper field operations are required. \ typedef bit = \UNIV :: bool set\ .. instantiation bit :: zero_neq_one begin definition zero_bit :: bit where \0 = Abs_bit False\ definition one_bit :: bit where \1 = Abs_bit True\ instance by standard (simp add: zero_bit_def one_bit_def Abs_bit_inject) end free_constructors case_bit for \0::bit\ | \1::bit\ proof - fix P :: bool fix a :: bit assume \a = 0 \ P\ and \a = 1 \ P\ then show P by (cases a) (auto simp add: zero_bit_def one_bit_def Abs_bit_inject) qed simp lemma bit_not_zero_iff [simp]: \a \ 0 \ a = 1\ for a :: bit by (cases a) simp_all lemma bit_not_one_imp [simp]: \a \ 1 \ a = 0\ for a :: bit by (cases a) simp_all instantiation bit :: semidom_modulo begin definition plus_bit :: \bit \ bit \ bit\ where \a + b = Abs_bit (Rep_bit a \ Rep_bit b)\ definition minus_bit :: \bit \ bit \ bit\ where [simp]: \minus_bit = plus\ definition times_bit :: \bit \ bit \ bit\ where \a * b = Abs_bit (Rep_bit a \ Rep_bit b)\ definition divide_bit :: \bit \ bit \ bit\ where [simp]: \divide_bit = times\ definition modulo_bit :: \bit \ bit \ bit\ where \a mod b = Abs_bit (Rep_bit a \ \ Rep_bit b)\ instance by standard (auto simp flip: Rep_bit_inject simp add: zero_bit_def one_bit_def plus_bit_def times_bit_def modulo_bit_def Abs_bit_inverse Rep_bit_inverse) end lemma bit_2_eq_0 [simp]: \2 = (0::bit)\ by (simp flip: one_add_one add: zero_bit_def plus_bit_def) instance bit :: semiring_parity apply standard apply (auto simp flip: Rep_bit_inject simp add: modulo_bit_def Abs_bit_inverse Rep_bit_inverse) apply (auto simp add: zero_bit_def one_bit_def Abs_bit_inverse Rep_bit_inverse) done lemma Abs_bit_eq_of_bool [code_abbrev]: \Abs_bit = of_bool\ by (simp add: fun_eq_iff zero_bit_def one_bit_def) lemma Rep_bit_eq_odd: \Rep_bit = odd\ proof - have \\ Rep_bit 0\ by (simp only: zero_bit_def) (subst Abs_bit_inverse, auto) then show ?thesis by (auto simp flip: Rep_bit_inject simp add: fun_eq_iff) qed lemma Rep_bit_iff_odd [code_abbrev]: \Rep_bit b \ odd b\ by (simp add: Rep_bit_eq_odd) lemma Not_Rep_bit_iff_even [code_abbrev]: \\ Rep_bit b \ even b\ by (simp add: Rep_bit_eq_odd) lemma Not_Not_Rep_bit [code_unfold]: \\ \ Rep_bit b \ Rep_bit b\ by simp code_datatype \0::bit\ \1::bit\ lemma Abs_bit_code [code]: \Abs_bit False = 0\ \Abs_bit True = 1\ by (simp_all add: Abs_bit_eq_of_bool) lemma Rep_bit_code [code]: \Rep_bit 0 \ False\ \Rep_bit 1 \ True\ by (simp_all add: Rep_bit_eq_odd) context zero_neq_one begin abbreviation of_bit :: \bit \ 'a\ where \of_bit b \ of_bool (odd b)\ end context begin qualified lemma bit_eq_iff: \a = b \ (even a \ even b)\ for a b :: bit by (cases a; cases b) simp_all end lemma modulo_bit_unfold [simp, code]: \a mod b = of_bool (odd a \ even b)\ for a b :: bit by (simp add: modulo_bit_def Abs_bit_eq_of_bool Rep_bit_eq_odd) lemma power_bit_unfold [simp, code]: \a ^ n = of_bool (odd a \ n = 0)\ for a :: bit by (cases a) simp_all instantiation bit :: semiring_bits begin definition bit_bit :: \bit \ nat \ bool\ where [simp]: \bit_bit b n \ odd b \ n = 0\ instance apply standard apply auto apply (metis bit.exhaust of_bool_eq(2)) done end instantiation bit :: semiring_bit_shifts begin definition push_bit_bit :: \nat \ bit \ bit\ where [simp]: \push_bit n b = of_bool (odd b \ n = 0)\ for b :: bit definition drop_bit_bit :: \nat \ bit \ bit\ where [simp]: \drop_bit_bit = push_bit\ definition take_bit_bit :: \nat \ bit \ bit\ where [simp]: \take_bit n b = of_bool (odd b \ n > 0)\ for b :: bit instance by standard simp_all end instantiation bit :: semiring_bit_operations begin definition and_bit :: \bit \ bit \ bit\ where [simp]: \b AND c = of_bool (odd b \ odd c)\ for b c :: bit definition or_bit :: \bit \ bit \ bit\ where [simp]: \b OR c = of_bool (odd b \ odd c)\ for b c :: bit definition xor_bit :: \bit \ bit \ bit\ where [simp]: \b XOR c = of_bool (odd b \ odd c)\ for b c :: bit definition mask_bit :: \nat \ bit\ where [simp]: \mask_bit n = of_bool (n > 0)\ instance by standard auto end lemma add_bit_eq_xor [simp, code]: \(+) = ((XOR) :: bit \ _)\ by (auto simp add: fun_eq_iff) lemma mult_bit_eq_and [simp, code]: \(*) = ((AND) :: bit \ _)\ by (simp add: fun_eq_iff) instantiation bit :: field begin definition uminus_bit :: \bit \ bit\ where [simp]: \uminus_bit = id\ definition inverse_bit :: \bit \ bit\ where [simp]: \inverse_bit = id\ instance by standard simp_all end instantiation bit :: ring_bit_operations begin definition not_bit :: \bit \ bit\ where [simp]: \NOT b = of_bool (even b)\ for b :: bit instance - by standard simp_all + by standard auto end lemma bit_numeral_even [simp]: "numeral (Num.Bit0 w) = (0 :: bit)" by (simp only: Z2.bit_eq_iff even_numeral) simp lemma bit_numeral_odd [simp]: "numeral (Num.Bit1 w) = (1 :: bit)" by (simp only: Z2.bit_eq_iff odd_numeral) simp end diff --git a/src/HOL/Parity.thy b/src/HOL/Parity.thy --- a/src/HOL/Parity.thy +++ b/src/HOL/Parity.thy @@ -1,1996 +1,1996 @@ (* 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) named_theorems bit_simps \Simplification rules for \<^const>\bit\\ lemma bit_exp_iff [bit_simps]: \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_simps]: \bit 1 n \ 1 \ 0 \ n = 0\ using bit_exp_iff [of 0 n] by simp lemma bit_2_iff [bit_simps]: \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_simps]: \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) + using \2 * 2 ^ n \ 0\ by simp (simp_all 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 bit_of_bool_iff [bit_simps]: \bit (of_bool b) n \ b \ n = 0\ by (simp add: bit_1_iff) 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_simps]: \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_simps]: \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_simps]: \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_simps]: \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 [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_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 [bit_simps]: \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_simps]: \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_not_int_iff': \bit (- k - 1) n \ \ bit k n\ for k :: int proof (induction n arbitrary: k) case 0 show ?case by simp next case (Suc n) have \- k - 1 = - (k + 2) + 1\ by simp also have \(- (k + 2) + 1) div 2 = - (k div 2) - 1\ proof (cases \even k\) case True then have \- k div 2 = - (k div 2)\ by rule (simp flip: mult_minus_right) with True show ?thesis by simp next case False have \4 = 2 * (2::int)\ by simp also have \2 * 2 div 2 = (2::int)\ by (simp only: nonzero_mult_div_cancel_left) finally have *: \4 div 2 = (2::int)\ . from False obtain l where k: \k = 2 * l + 1\ .. then have \- k - 2 = 2 * - (l + 2) + 1\ by simp then have \(- k - 2) div 2 + 1 = - (k div 2) - 1\ by (simp flip: mult_minus_right add: *) (simp add: k) with False show ?thesis by simp qed finally have \(- k - 1) div 2 = - (k div 2) - 1\ . with Suc show ?case by (simp add: bit_Suc) qed lemma bit_minus_int_iff [bit_simps]: \bit (- k) n \ \ bit (k - 1) n\ for k :: int using bit_not_int_iff' [of \k - 1\] by simp lemma bit_nat_iff [bit_simps]: \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 (simp add: bit_simps) 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 (subst (asm) 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/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2692 +1,2736 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all -lemma split_of_bool_asm: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" +lemma split_of_bool_asm [split]: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all +lemma of_bool_eq_0_iff [simp]: + \of_bool P = 0 \ \ P\ + by simp + +lemma of_bool_eq_1_iff [simp]: + \of_bool P = 1 \ P\ + by simp + end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" using that by rule (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. +lemma of_bool_not_iff [simp]: + \of_bool (\ P) = 1 - of_bool P\ + by simp + lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" by (rule dvd_div_eq_0_iff) (insert assms, auto) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ begin lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right end class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp lemma normalize_mult_unit_left [simp]: assumes "a dvd 1" shows "normalize (a * b) = normalize b" proof (cases "b = 0") case False have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" using assms by (subst unit_factor_mult_unit_left) auto also have "\ = a * b" by simp also have "b = unit_factor b * normalize b" by simp hence "a * b = a * unit_factor b * normalize b" by (simp only: mult_ac) finally show ?thesis using assms False by auto qed auto lemma normalize_mult_unit_right [simp]: assumes "b dvd 1" shows "normalize (a * b) = normalize a" using assms by (subst mult.commute) auto lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") case False have "normalize a = normalize (unit_factor a * normalize a)" by simp also from False have "\ = normalize (normalize a)" by (subst normalize_mult_unit_left) auto finally show ?thesis .. qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) end class normalization_semidom_multiplicative = normalization_semidom + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" begin lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False have "unit_factor (a * b) * normalize (a * b) = a * b" by (rule unit_factor_mult_normalize) then have "normalize (a * b) = a * b div unit_factor (a * b)" by simp also have "\ = a * b div unit_factor (b * a)" by (simp add: ac_simps) also have "\ = a * b div unit_factor b div unit_factor a" using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) also have "\ = a * (b div unit_factor b) div unit_factor a" using False by (subst unit_div_mult_swap) simp_all also have "\ = normalize a * normalize b" using False by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) finally show ?thesis . qed lemma dvd_unit_factor_div: assumes "b dvd a" shows "unit_factor (a div b) = unit_factor a div unit_factor b" proof - from assms have "a = a div b * b" by simp then have "unit_factor a = unit_factor (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: unit_factor_mult) qed lemma dvd_normalize_div: assumes "b dvd a" shows "normalize (a div b) = normalize a div normalize b" proof - from assms have "a = a div b * b" by simp then have "normalize a = normalize (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: normalize_mult) qed end text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: assumes "0 < a * b" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b]) qed (auto simp add: le_less not_less) lemma zero_less_mult_pos2: assumes "0 < b * a" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b]) qed (auto simp add: le_less not_less) text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" "c < d" "0 < b" "0 \ c" shows "a * c < b * d" proof (cases "c = 0") case True with assms show ?thesis by simp next case False with assms have "a*c < b*c" by (simp add: mult_strict_right_mono [OF \a < b\]) also have "\ < b*d" by (simp add: assms mult_strict_left_mono) finally show ?thesis . qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" proof - have "a * c < b * c" by (simp add: assms mult_strict_right_mono) also have "... \ b * d" by (intro mult_left_mono) (use assms in auto) finally show ?thesis . qed lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" proof - have "a * c \ b * c" by (simp add: assms mult_right_mono) also have "... < b * d" by (intro mult_strict_left_mono) (use assms in auto) finally show ?thesis . qed end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" by (auto dest: mult_left_mono [of _ _ "- c"]) lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" by (auto dest: mult_right_mono [of _ _ "- c"]) lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg) qed qed auto lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg) qed qed auto text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one by standard (insert zero_less_one, blast) subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) +lemma of_bool_less_eq_iff [simp]: + \of_bool P \ of_bool Q \ (P \ Q)\ + by auto + +lemma of_bool_less_iff [simp]: + \of_bool P < of_bool Q \ \ P \ Q\ + by auto + lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b proof (rule ccontr, simp add: not_less) assume "b \ a" with that show False by (simp add: ) qed qed +lemma zero_less_eq_of_bool [simp]: + \0 \ of_bool P\ + by simp + +lemma zero_less_of_bool_iff [simp]: + \0 < of_bool P \ P\ + by simp + +lemma of_bool_less_eq_one [simp]: + \of_bool P \ 1\ + by simp + +lemma of_bool_less_one_iff [simp]: + \of_bool P < 1 \ \ P\ + by simp + +lemma of_bool_or_iff [simp]: + \of_bool (P \ Q) = max (of_bool P) (of_bool Q)\ + by (simp add: max_def) + text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: assumes "i + k \ n" shows "i \ n - k" proof - have "n - (i + k) + i + k = n" by (simp add: assms add.assoc) with assms add_implies_diff have "i + k \ n - k + k" by fastforce then show ?thesis by simp qed lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + i + k = n" using 1 by (simp add: add.assoc) moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 add_le_imp_le_diff [of "n-k" k "j + k"] by (simp add: add.commute diff_diff_add) qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) +lemma abs_bool_eq [simp]: + \\of_bool P\ = of_bool P\ + by simp + lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end