diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,649 +1,127 @@ (* Title: HOL/Divides.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge *) -section \More on quotient and remainder\ +section \Lemmas of doubtful value\ theory Divides imports Parity begin -subsection \More on division\ - -subsubsection \Monotonicity in the First Argument (Dividend)\ - -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)" - using unique_quotient_lemma[where b = "-b" and r = "-r'" and r'="-r"] by auto - -lemma zdiv_mono1: - \a div b \ a' div b\ - if \a \ a'\ \0 < b\ - for a b b' :: int -proof (rule unique_quotient_lemma) - show "b * (a div b) + a mod b \ b * (a' div b) + a' mod b" - using \a \ a'\ by auto -qed (use that 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) - - -subsubsection \Quotients of Signs\ - -lemma div_eq_minus1: "0 < b \ - 1 div b = - 1" for b :: int - by (simp add: divide_int_def) - -lemma zmod_minus1: "0 < b \ - 1 mod b = b - 1" for b :: int - by (auto simp add: modulo_int_def) - -lemma minus_mod_int_eq: - \- k mod l = l - 1 - (k - 1) mod l\ if \l \ 0\ for k l :: int -proof (cases \l = 0\) - case True - then show ?thesis - by simp -next - case False - with that have \l > 0\ - by simp - then show ?thesis - proof (cases \l dvd k\) - case True - then obtain j where \k = l * j\ .. - moreover have \(l * j mod l - 1) mod l = l - 1\ - using \l > 0\ by (simp add: zmod_minus1) - then have \(l * j - 1) mod l = l - 1\ - by (simp only: mod_simps) - ultimately show ?thesis - by simp - next - case False - moreover have 1: \0 < k mod l\ - using \0 < l\ False le_less by fastforce - moreover have 2: \k mod l < 1 + l\ - using \0 < l\ pos_mod_bound[of l k] by linarith - from 1 2 \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ - by (simp add: zmod_trivial_iff) - ultimately show ?thesis - by (simp only: zmod_zminus1_eq_if) - (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) - qed -qed - -lemma div_neg_pos_less0: - 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) - -lemma sgn_div_eq_sgn_mult: - \sgn (k div l) = of_bool (k div l \ 0) * sgn (k * l)\ - for k l :: int -proof (cases \k div l = 0\) - case True - then show ?thesis - by simp -next - case False - have \0 \ \k\ div \l\\ - by (cases \l = 0\) (simp_all add: pos_imp_zdiv_nonneg_iff) - then have \\k\ div \l\ \ 0 \ 0 < \k\ div \l\\ - by (simp add: less_le) - also have \\ \ \k\ \ \l\\ - using False nonneg1_imp_zdiv_pos_iff by auto - finally have *: \\k\ div \l\ \ 0 \ \l\ \ \k\\ . - show ?thesis - using \0 \ \k\ div \l\\ False - by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l] - sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp) -qed - -lemma - fixes a b q r :: int - assumes \a = b * q + r\ \0 \ r\ \r < b\ - shows int_div_pos_eq: - \a div b = q\ (is ?Q) - and int_mod_pos_eq: - \a mod b = r\ (is ?R) -proof - - from assms have \(a div b, a mod b) = (q, r)\ - by (cases b q r a rule: euclidean_relation_intI) - (auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le) - then show ?Q and ?R - by simp_all -qed - -lemma int_div_neg_eq: - \a div b = q\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int - using that int_div_pos_eq [of a \- b\ \- q\ \- r\] by simp_all - -lemma int_mod_neg_eq: - \a mod b = r\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int - using that int_div_neg_eq [of a b q r] by simp - - -subsubsection \Further properties\ - -lemma div_int_pos_iff: - "k div l \ 0 \ k = 0 \ l = 0 \ k \ 0 \ l \ 0 - \ k < 0 \ l < 0" - for k l :: int -proof (cases "k = 0 \ l = 0") - case False - then have *: "k \ 0" "l \ 0" - by auto - then have "0 \ k div l \ \ k < 0 \ 0 \ l" - by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) - then show ?thesis - using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) -qed auto - -lemma mod_int_pos_iff: - "k mod l \ 0 \ l dvd k \ l = 0 \ k \ 0 \ l > 0" - for k l :: int -proof (cases "l > 0") - case False - then show ?thesis - by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \auto simp add: le_less not_less\) -qed auto - -text \Simplify expressions in which div and mod combine numerical constants\ - -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 exI) - -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 - proof - fix q - assume "y = x + n * q" - then have "x + n * q = y + n * 0" - by simp - then show "\q1 q2. x + n * q1 = y + n * q2" - by blast - qed} - moreover - {assume xy: "y \ x" - from nat_mod_eq_lemma[OF H xy] have ?rhs - proof - fix q - assume "x = y + n * q" - then have "x + n * 0 = y + n * q" - by simp - then show "\q1 q2. x + n * q1 = y + n * q2" - by blast - qed} - 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 - - -subsubsection \Computing \div\ and \mod\ with shifting\ - -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 - -text\computing div by shifting\ - -lemma pos_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = b div a\ (is ?Q) - and pos_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\ (is ?R) - if \0 \ a\ for a b :: int -proof - - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\ - proof (cases \2 * a\ \b div a\ \1 + 2 * (b mod a)\ \1 + 2 * b\ rule: euclidean_relation_intI) - case by0 - then show ?case - by simp - next - case divides - have \even (2 * a)\ - by simp - then have \even (1 + 2 * b)\ - using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) - then show ?case - by simp - next - case euclidean_relation - with that have \a > 0\ - by simp - moreover have \b mod a < a\ - using \a > 0\ by simp - then have \1 + 2 * (b mod a) < 2 * a\ - by simp - moreover have \2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\ - by (simp only: algebra_simps) - moreover have \0 \ 2 * (b mod a)\ - using \a > 0\ by simp - ultimately show ?case - by (simp add: algebra_simps) - qed - then show ?Q and ?R - by simp_all -qed - -lemma neg_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = (b + 1) div a\ (is ?Q) - and neg_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\ (is ?R) - if \a \ 0\ for a b :: int -proof - - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\ - proof (cases \2 * a\ \(b + 1) div a\ \2 * ((b + 1) mod a) - 1\ \1 + 2 * b\ rule: euclidean_relation_intI) - case by0 - then show ?case - by simp - next - case divides - have \even (2 * a)\ - by simp - then have \even (1 + 2 * b)\ - using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) - then show ?case - by simp - next - case euclidean_relation - with that have \a < 0\ - by simp - moreover have \(b + 1) mod a > a\ - using \a < 0\ by simp - then have \2 * ((b + 1) mod a) > 1 + 2 * a\ - by simp - moreover have \((1 + b) mod a) \ 0\ - using \a < 0\ by simp - then have \2 * ((1 + b) mod a) \ 0\ - by simp - moreover have \2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) = - 2 * ((1 + b) div a * a + (1 + b) mod a)\ - by (simp only: algebra_simps) - ultimately show ?case - by (simp add: algebra_simps sgn_mult abs_mult) - qed - then show ?Q and ?R - by simp_all -qed - -lemma zdiv_numeral_Bit0 [simp]: - "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = - numeral v div (numeral w :: int)" - unfolding numeral.simps unfolding mult_2 [symmetric] - by (rule div_mult_mult1, simp) - -lemma zdiv_numeral_Bit1 [simp]: - "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = - (numeral v div (numeral w :: int))" - unfolding numeral.simps - unfolding mult_2 [symmetric] add.commute [of _ 1] - by (rule pos_zdiv_mult_2, simp) - -lemma zmod_numeral_Bit0 [simp]: - "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = - (2::int) * (numeral v mod numeral w)" - unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] - unfolding mult_2 [symmetric] by (rule mod_mult_mult1) - -lemma zmod_numeral_Bit1 [simp]: - "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = - 2 * (numeral v mod numeral w) + (1::int)" - unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] - unfolding mult_2 [symmetric] add.commute [of _ 1] - by (rule pos_zmod_mult_2, simp) - - -code_identifier - code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith - - -subsection \Lemmas of doubtful value\ - 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" 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" using mod_less by (auto simp add: mod_w) with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp have "2 * (a div (2 * b)) = a div b - w" by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) with \w = 1\ have div: "2 * (a div (2 * b)) = a div b - 1" by simp then show ?P and ?Q by (simp_all add: div mod add_implies_diff [symmetric]) qed lemma divmod_digit_0: 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 end hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq instance nat :: unique_euclidean_semiring_numeral by standard (auto simp add: div_greater_zero_iff div_mult2_eq mod_mult2_eq) instance int :: unique_euclidean_semiring_numeral by standard (auto intro: zmod_le_nonneg_dividend simp add: pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \ m < n" for m n :: nat by (rule le_div_geq) (use that in \simp_all add: not_less\) lemma mod_geq: "m mod n = (m - n) mod n" if "\ m < n" for m n :: nat by (rule le_mod_geq) (use that in \simp add: not_less\) lemma mod_eq_0D: "\q. m = d * q" if "m mod d = 0" for m d :: nat using that by (auto simp add: mod_eq_0_iff_dvd) lemma pos_mod_conj: "0 < b \ 0 \ a mod b \ a mod b < b" for a b :: int by simp lemma neg_mod_conj: "b < 0 \ a mod b \ 0 \ b < a mod b" for a b :: int by simp lemma zmod_eq_0_iff: "m mod d = 0 \ (\q. m = d * q)" for m d :: int by (auto simp add: mod_eq_0_iff_dvd) (* REVISIT: should this be generalized to all semiring_div types? *) lemma zmod_eq_0D [dest!]: "\q. m = d * q" if "m mod d = 0" for m d :: int using that by auto lemma div_positive_int: "k div l > 0" if "k \ l" and "l > 0" for k l :: int using that by (simp add: nonneg1_imp_zdiv_pos_iff) +code_identifier + code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith + end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,2818 +1,3319 @@ (* 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 + + +class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" - assumes mod_size_less: + 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" + 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: +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] + with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (use assms in simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: - assumes "a dvd b" "b \ 0" + assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto simp: size_mult_mono) lemma dvd_proper_imp_size_less: - assumes "a dvd b" "\ b dvd a" "b \ 0" + assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule iffI; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed - + end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") auto lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ -lemma power_mod [mod_simps]: +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 + +class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: \euclidean_size (a * b) = euclidean_size a * euclidean_size b\ fixes division_segment :: \'a \ 'a\ assumes is_unit_division_segment [simp]: \is_unit (division_segment a)\ and division_segment_mult: \a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b\ and division_segment_mod: \b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b\ assumes div_bounded: \b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q\ begin lemma division_segment_not_0 [simp]: \division_segment a \ 0\ using is_unit_division_segment [of a] is_unitE [of \division_segment a\] by blast lemma euclidean_relationI [case_names by0 divides euclidean_relation]: \(a div b, a mod b) = (q, r)\ if by0: \b = 0 \ q = 0 \ r = a\ and divides: \b \ 0 \ b dvd a \ r = 0 \ a = q * b\ and euclidean_relation: \b \ 0 \ \ b dvd a \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ a = q * b + r\ proof (cases \b = 0\) case True with by0 show ?thesis by simp next case False show ?thesis proof (cases \b dvd a\) case True with \b \ 0\ divides show ?thesis by simp next case False with \b \ 0\ euclidean_relation have \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ \a = q * b + r\ by simp_all from \b \ 0\ \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ have \(q * b + r) div b = q\ by (rule div_bounded) with \a = q * b + r\ have \q = a div b\ by simp from \a = q * b + r\ have \a div b * b + a mod b = q * b + r\ by (simp add: div_mult_mod_eq) with \q = a div b\ have \q * b + a mod b = q * b + r\ by simp then have \r = a mod b\ by simp with \q = a div b\ show ?thesis by simp qed qed subclass euclidean_semiring_cancel proof fix a b c assume \b \ 0\ have \((a + c * b) div b, (a + c * b) mod b) = (c + a div b, a mod b)\ proof (cases b \c + a div b\ \a mod b\ \a + c * b\ rule: euclidean_relationI) case by0 with \b \ 0\ show ?case by simp next case divides then show ?case by (simp add: algebra_simps dvd_add_left_iff) next case euclidean_relation then have \\ b dvd a\ by (simp add: dvd_add_left_iff) have \a mod b + (b * c + b * (a div b)) = b * c + ((a div b) * b + a mod b)\ by (simp add: ac_simps) with \b \ 0\ have *: \a mod b + (b * c + b * (a div b)) = b * c + a\ by (simp add: div_mult_mod_eq) from \\ b dvd a\ euclidean_relation show ?case by (simp_all add: algebra_simps division_segment_mod mod_size_less *) qed then show \(a + c * b) div b = c + a div b\ by simp next fix a b c assume \c \ 0\ have \((c * a) div (c * b), (c * a) mod (c * b)) = (a div b, c * (a mod b))\ proof (cases \c * b\ \a div b\ \c * (a mod b)\ \c * a\ rule: euclidean_relationI) case by0 with \c \ 0\ show ?case by simp next case divides then show ?case by (auto simp add: algebra_simps) next case euclidean_relation then have \b \ 0\ \a mod b \ 0\ by (simp_all add: mod_eq_0_iff_dvd) have \c * (a mod b) + b * (c * (a div b)) = c * ((a div b) * b + a mod b)\ by (simp add: algebra_simps) with \b \ 0\ have *: \c * (a mod b) + b * (c * (a div b)) = c * a\ by (simp add: div_mult_mod_eq) from \b \ 0\ \c \ 0\ have \euclidean_size c * euclidean_size (a mod b) < euclidean_size c * euclidean_size b\ using mod_size_less [of b a] by simp with euclidean_relation \b \ 0\ \a mod b \ 0\ show ?case by (simp add: algebra_simps division_segment_mult division_segment_mod euclidean_size_mult *) qed then show \(c * a) div (c * b) = a div b\ by simp qed lemma div_eq_0_iff: \a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0\ (is "_ \ ?P") if \division_segment a = division_segment b\ proof (cases \a = 0 \ b = 0\) case True then show ?thesis by auto next case False then have \a \ 0\ \b \ 0\ by simp_all have \a div b = 0 \ euclidean_size a < euclidean_size b\ proof assume \a div b = 0\ then have \a mod b = a\ using div_mult_mod_eq [of a b] by simp with \b \ 0\ mod_size_less [of b a] show \euclidean_size a < euclidean_size b\ by simp next assume \euclidean_size a < euclidean_size b\ have \(a div b, a mod b) = (0, a)\ proof (cases b 0 a a rule: euclidean_relationI) case by0 show ?case by simp next case divides with \euclidean_size a < euclidean_size b\ show ?case using dvd_imp_size_le [of b a] \a \ 0\ by simp next case euclidean_relation with \euclidean_size a < euclidean_size b\ that show ?case by simp qed then show \a div b = 0\ by simp qed with \b \ 0\ show ?thesis by simp qed lemma div_mult1_eq: \(a * b) div c = a * (b div c) + a * (b mod c) div c\ proof - have *: \(a * b) mod c + (a * (c * (b div c)) + c * (a * (b mod c) div c)) = a * b\ (is \?A + (?B + ?C) = _\) proof - have \?A = a * (b mod c) mod c\ by (simp add: mod_mult_right_eq) then have \?C + ?A = a * (b mod c)\ by (simp add: mult_div_mod_eq) then have \?B + (?C + ?A) = a * (c * (b div c) + (b mod c))\ by (simp add: algebra_simps) also have \\ = a * b\ by (simp add: mult_div_mod_eq) finally show ?thesis by (simp add: algebra_simps) qed have \((a * b) div c, (a * b) mod c) = (a * (b div c) + a * (b mod c) div c, (a * b) mod c)\ proof (cases c \a * (b div c) + a * (b mod c) div c\ \(a * b) mod c\ \a * b\ rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed lemma div_add1_eq: \(a + b) div c = a div c + b div c + (a mod c + b mod c) div c\ proof - have *: \(a + b) mod c + (c * (a div c) + (c * (b div c) + c * ((a mod c + b mod c) div c))) = a + b\ (is \?A + (?B + (?C + ?D)) = _\) proof - have \?A + (?B + (?C + ?D)) = ?A + ?D + (?B + ?C)\ by (simp add: ac_simps) also have \?A + ?D = (a mod c + b mod c) mod c + ?D\ by (simp add: mod_add_eq) also have \\ = a mod c + b mod c\ by (simp add: mod_mult_div_eq) finally have \?A + (?B + (?C + ?D)) = (a mod c + ?B) + (b mod c + ?C)\ by (simp add: ac_simps) then show ?thesis by (simp add: mod_mult_div_eq) qed have \((a + b) div c, (a + b) mod c) = (a div c + b div c + (a mod c + b mod c) div c, (a + b) mod c)\ proof (cases c \a div c + b div c + (a mod c + b mod c) div c\ \(a + b) mod c\ \a + b\ rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin - + subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: \nat \ nat\ where [simp]: \normalize = (id :: nat \ nat)\ definition unit_factor_nat :: \nat \ nat\ where \unit_factor n = of_bool (n > 0)\ for n :: nat lemma unit_factor_simps [simp]: \unit_factor 0 = (0::nat)\ \unit_factor (Suc n) = 1\ by (simp_all add: unit_factor_nat_def) definition divide_nat :: \nat \ nat \ nat\ where \m div n = (if n = 0 then 0 else Max {k. k * n \ m})\ for m n :: nat instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: \nat \ nat\ where [simp]: \euclidean_size_nat = id\ definition division_segment_nat :: \nat \ nat\ where [simp]: \division_segment n = 1\ for n :: nat definition modulo_nat :: \nat \ nat \ nat\ where \m mod n = m - (m div n * n)\ for m n :: nat instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) - assume "n \ 0" + assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end lemma euclidean_relation_natI [case_names by0 divides euclidean_relation]: \(m div n, m mod n) = (q, r)\ if by0: \n = 0 \ q = 0 \ r = m\ and divides: \n > 0 \ n dvd m \ r = 0 \ m = q * n\ and euclidean_relation: \n > 0 \ \ n dvd m \ r < n \ m = q * n + r\ for m n q r :: nat by (rule euclidean_relationI) (use that in simp_all) lemma div_nat_eqI: \m div n = q\ if \n * q \ m\ and \m < n * Suc q\ for m n q :: nat proof - have \(m div n, m mod n) = (q, m - n * q)\ proof (cases n q \m - n * q\ m rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from \n dvd m\ obtain s where \m = n * s\ .. with \n > 0\ that have \s < Suc q\ by (simp only: mult_less_cancel1) with \m = n * s\ \n > 0\ that have \q = s\ by simp with \m = n * s\ show ?case by (simp add: ac_simps) next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed lemma mod_nat_eqI: \m mod n = r\ if \r < n\ and \r \ m\ and \n dvd m - r\ for m n r :: nat proof - have \(m div n, m mod n) = ((m - r) div n, r)\ proof (cases n \(m - r) div n\ r m rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from that dvd_minus_add [of r \m\ 1 n] have \n dvd m + (n - r)\ by simp with divides have \n dvd n - r\ by (simp add: dvd_add_right_iff) then have \n \ n - r\ by (rule dvd_imp_le) (use \r < n\ in simp) with \n > 0\ have \r = 0\ by simp with \n > 0\ that show ?case by simp next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_nat_eqI mod_nat_eqI) - + lemma split_div: \P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))\ (is ?div) and split_mod: \Q (m mod n) \ (n = 0 \ Q m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ Q j))\ (is ?mod) for m n :: nat proof - have *: \R (m div n) (m mod n) \ (n = 0 \ R 0 m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ R i j))\ for R by (cases \n = 0\) auto from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed declare split_div [of _ _ \numeral n\, linarith_split] for n declare split_mod [of _ _ \numeral n\, linarith_split] for n lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis - by auto + by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all lemma div_mult2_eq: \m div (n * q) = (m div n) div q\ (is ?Q) and mod_mult2_eq: \m mod (n * q) = n * (m div n mod q) + m mod n\ (is ?R) for m n q :: nat proof - have \(m div (n * q), m mod (n * q)) = ((m div n) div q, n * (m div n mod q) + m mod n)\ proof (cases \n * q\ \(m div n) div q\ \n * (m div n mod q) + m mod n\ m rule: euclidean_relation_natI) case by0 then show ?case by auto next case divides from \n * q dvd m\ obtain t where \m = n * q * t\ .. with \n * q > 0\ show ?case by (simp add: algebra_simps) next case euclidean_relation then have \n > 0\ \q > 0\ by simp_all from \n > 0\ have \m mod n < n\ by (rule mod_less_divisor) from \q > 0\ have \m div n mod q < q\ by (rule mod_less_divisor) then obtain s where \q = Suc (m div n mod q + s)\ by (blast dest: less_imp_Suc_add) moreover have \m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)\ using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately have \m mod n + n * (m div n mod q) < n * q\ by simp then show ?case by (simp add: algebra_simps flip: add_mult_distrib2) qed then show ?Q and ?R by simp_all qed lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed lemma div_less_iff_less_mult: \m div q < n \ m < n * q\ (is \?P \ ?Q\) if \q > 0\ for m n q :: nat proof assume ?Q then show ?P by (rule less_mult_imp_div_less) next assume ?P then obtain h where \n = Suc (m div q + h)\ using less_natE by blast moreover have \m < m + (Suc h * q - m mod q)\ using that by (simp add: trans_less_add1) ultimately show ?Q by (simp add: algebra_simps flip: minus_mod_eq_mult_div) qed lemma less_eq_div_iff_mult_less_eq: \m \ n div q \ m * q \ n\ if \q > 0\ for m n q :: nat using div_less_iff_less_mult [of q n m] that by auto text \A fact for the mutilated chess board\ lemma 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) + by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma funpow_mod_eq: \<^marker>\contributor \Lars Noschinski\\ \(f ^^ (m mod n)) x = (f ^^ m) x\ if \(f ^^ n) x = x\ proof - have \(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\ by simp also have \\ = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\ by (simp only: funpow_add funpow_mult ac_simps) simp also have \((f ^^ n) ^^ q) x = x\ for q by (induction q) (use \(f ^^ n) x = x\ in simp_all) finally show ?thesis by simp qed +lemma mod_eq_dvd_iff_nat: + \m mod q = n mod q \ q dvd m - n\ (is \?P \ ?Q\) + if \m \ n\ for m n q :: nat +proof + assume ?Q + then obtain s where \m - n = q * s\ .. + with that have \m = q * s + n\ + by simp + then show ?P + by simp +next + assume ?P + have \m - n = m div q * q + m mod q - (n div q * q + n mod q)\ + by simp + also have \\ = q * (m div q - n div q)\ + by (simp only: algebra_simps \?P\) + finally show ?Q .. +qed + +lemma mod_eq_nat1E: + fixes m n q :: nat + assumes "m mod q = n mod q" and "m \ n" + obtains s where "m = n + q * s" +proof - + from assms have "q dvd m - n" + by (simp add: mod_eq_dvd_iff_nat) + then obtain s where "m - n = q * s" .. + with \m \ n\ have "m = n + q * s" + by simp + with that show thesis . +qed + +lemma mod_eq_nat2E: + fixes m n q :: nat + assumes "m mod q = n mod q" and "n \ m" + obtains s where "n = m + q * s" + using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) + +lemma nat_mod_eq_iff: + "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" (is "?lhs = ?rhs") +proof + assume H: "x mod n = y mod n" + { assume xy: "x \ y" + from H have th: "y mod n = x mod n" by simp + from mod_eq_nat1E [OF th xy] obtain q where "y = x + n * q" . + then have "x + n * q = y + n * 0" + by simp + then have "\q1 q2. x + n * q1 = y + n * q2" + by blast + } + moreover + { assume xy: "y \ x" + from mod_eq_nat1E [OF H xy] obtain q where "x = y + n * q" . + then have "x + n * 0 = y + n * q" + by simp + then have "\q1 q2. x + n * q1 = y + n * q2" + by blast + } + ultimately show ?rhs using linear[of x y] by blast +next + assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast + hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp + thus ?lhs by simp +qed + + subsection \Elementary euclidean division on \<^typ>\int\\ subsubsection \Basic instantiation\ instantiation int :: "{normalization_semidom, idom_modulo}" begin definition normalize_int :: \int \ int\ where [simp]: \normalize = (abs :: int \ int)\ definition unit_factor_int :: \int \ int\ where [simp]: \unit_factor = (sgn :: int \ int)\ definition divide_int :: \int \ int \ int\ where \k div l = (sgn k * sgn l * int (nat \k\ div nat \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k))\ lemma divide_int_unfold: \(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n) - of_bool ((k = 0 \ m = 0) \ l \ 0 \ n \ 0 \ sgn k \ sgn l \ \ n dvd m))\ by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps) definition modulo_int :: \int \ int \ int\ where \k mod l = sgn k * int (nat \k\ mod nat \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ lemma modulo_int_unfold: \(sgn k * int m) mod (sgn l * int n) = sgn k * int (m mod (of_bool (l \ 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \ m = 0) \ sgn k \ sgn l \ \ n dvd m)\ by (auto simp add: modulo_int_def sgn_mult abs_mult) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" - obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" + obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) next fix k l :: int - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" + obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end subsubsection \Algebraic foundations\ lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all subsubsection \Basic conversions\ lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (auto simp add: divide_int_def) lemma div_eq_div_abs: \k div l = sgn k * sgn l * (\k\ div \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: divide_int_def [of k l] div_abs_eq_div_nat) lemma div_abs_eq: \\k\ div \l\ = sgn k * sgn l * (k div l + of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (simp add: div_eq_div_abs [of k l] ac_simps) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma mod_eq_mod_abs: \k mod l = sgn k * (\k\ mod \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat) lemma mod_abs_eq: \\k\ mod \l\ = sgn k * (k mod l - l * of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (auto simp: mod_eq_mod_abs [of k l]) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" using assms by (simp add: sgn_mult abs_mult sgn_0_0 divide_int_def [of "sgn v * \k\" "sgn v * \l\"] flip: div_abs_eq_div_nat) lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" using assms by (auto simp add: div_abs_eq) lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" using assms by (auto simp add: div_abs_eq ac_simps) lemma div_noneq_sgn_abs: fixes k l :: int assumes "l \ 0" assumes "sgn k \ sgn l" shows "k div l = - (\k\ div \l\) - of_bool (\ l dvd k)" using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp) subsubsection \Euclidean division\ instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" + obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) (simp add: sgn_0_0) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" + obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: * q l divide_int_unfold) (auto simp add: sgn_mult ac_simps) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) end lemma euclidean_relation_intI [case_names by0 divides euclidean_relation]: \(k div l, k mod l) = (q, r)\ if by0': \l = 0 \ q = 0 \ r = k\ and divides': \l \ 0 \ l dvd k \ r = 0 \ k = q * l\ and euclidean_relation': \l \ 0 \ \ l dvd k \ sgn r = sgn l \ \r\ < \l\ \ k = q * l + r\ for k l :: int proof (cases l q r k rule: euclidean_relationI) case by0 then show ?case by (rule by0') next case divides then show ?case by (rule divides') next case euclidean_relation with euclidean_relation' have \sgn r = sgn l\ \\r\ < \l\\ \k = q * l + r\ by simp_all from \sgn r = sgn l\ \l \ 0\ have \division_segment r = division_segment l\ by (simp add: division_segment_int_def sgn_if split: if_splits) with \\r\ < \l\\ \k = q * l + r\ show ?case by simp qed subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ proof (cases \m = 0 \ n = 0\) case True then show ?thesis by auto next case False then have \m > 0\ \n > 0\ by simp_all show ?thesis proof (cases \of_nat m * of_nat n dvd a\) case True then obtain b where \a = (of_nat m * of_nat n) * b\ .. then have \a = of_nat m * (of_nat n * b)\ by (simp add: ac_simps) then show ?thesis by simp next case False define q where \q = a div (of_nat m * of_nat n)\ define r where \r = a mod (of_nat m * of_nat n)\ from \m > 0\ \n > 0\ \\ of_nat m * of_nat n dvd a\ r_def have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod) with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with \m > 0\ \n > 0\ r_def have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with \m > 0\ \n > 0\ q_def have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp moreover have \a = q * (of_nat m * of_nat n) + r\ by (simp add: q_def r_def div_mult_mod_eq) ultimately show \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ using q_def [symmetric] div_plus_div_distrib_dvd_right [of \of_nat m\ \q * (of_nat m * of_nat n)\ r] by (simp add: ac_simps) qed qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np) subsection \More on euclidean division on \<^typ>\int\\ subsubsection \Trivial reduction steps\ lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) -lemma div_pos_neg_trivial: - \k div l = - 1\ if \0 < k\ and \k + l \ 0\ for k l :: int -proof (cases \l dvd k\) - case True - then obtain j where \k = l * j\ .. +lemma + div_pos_neg_trivial: \k div l = - 1\ (is ?Q) + and mod_pos_neg_trivial: \k mod l = k + l\ (is ?R) + if \0 < k\ and \k + l \ 0\ for k l :: int +proof - from that have \l < 0\ by simp - with \k = l * j\ \0 < k\ have \j \ - 1\ - by (simp add: zero_less_mult_iff) - moreover from \k + l \ 0\ \k = l * j\ have \l * (j + 1) \ 0\ - by (simp add: algebra_simps) - with \l < 0\ have \- 1 \ j\ - by (simp add: mult_le_0_iff) - ultimately have \j = - 1\ - by (rule order.antisym) - with \k = l * j\ \l < 0\ show ?thesis - by (simp add: dvd_neg_div) -next - case False - have \k + l < 0\ - proof (rule ccontr) - assume \\ k + l < 0\ - with \k + l \ 0\ have \k + l = 0\ + have \(k div l, k mod l) = (- 1, k + l)\ + proof (cases l \- 1 :: int\ \k + l\ k rule: euclidean_relation_intI) + case by0 + with \l < 0\ show ?case by simp - then have \k = - l\ + next + case divides + from \l dvd k\ obtain j where \k = l * j\ .. + with \l < 0\ \0 < k\ have \j < 0\ + by (simp add: zero_less_mult_iff) + moreover from \k + l \ 0\ \k = l * j\ have \l * (j + 1) \ 0\ + by (simp add: algebra_simps) + with \l < 0\ have \j + 1 \ 0\ + by (simp add: mult_le_0_iff) + with \j < 0\ have \j = - 1\ by simp - then have \l dvd k\ + with \k = l * j\ show ?case by simp - with \\ l dvd k\ show False .. + next + case euclidean_relation + with \k + l \ 0\ have \k + l < 0\ + by (auto simp add: less_le add_eq_0_iff) + with \0 < k\ show ?case + by simp qed - with that \\ l dvd k\ show ?thesis - by (simp add: div_eq_div_abs [of k l]) -qed - -lemma mod_pos_neg_trivial: - \k mod l = k + l\ if \0 < k\ and \k + l \ 0\ for k l :: int -proof - - from that have \k mod l = k div l * l + k mod l + l\ - by (simp add: div_pos_neg_trivial) - then show ?thesis by simp + then show ?Q and ?R + by simp_all qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ +subsubsection \More uniqueness rules\ + +lemma + fixes a b q r :: int + assumes \a = b * q + r\ \0 \ r\ \r < b\ + shows int_div_pos_eq: + \a div b = q\ (is ?Q) + and int_mod_pos_eq: + \a mod b = r\ (is ?R) +proof - + from assms have \(a div b, a mod b) = (q, r)\ + by (cases b q r a rule: euclidean_relation_intI) + (auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le) + then show ?Q and ?R + by simp_all +qed + +lemma int_div_neg_eq: + \a div b = q\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int + using that int_div_pos_eq [of a \- b\ \- q\ \- r\] by simp_all + +lemma int_mod_neg_eq: + \a mod b = r\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int + using that int_div_neg_eq [of a b q r] by simp + + subsubsection \Laws for unary minus\ lemma zmod_zminus1_not_zero: fixes k l :: int shows "- k mod l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zmod_zminus2_not_zero: fixes k l :: int shows "k mod - l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zdiv_zminus1_eq_if: \(- a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that sgn_not_eq_imp [of b \- a\] by (cases \a = 0\) (auto simp add: div_eq_div_abs [of \- a\ b] div_eq_div_abs [of a b] sgn_eq_0_iff) lemma zdiv_zminus2_eq_if: \a div (- b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that by (auto simp add: zdiv_zminus1_eq_if div_minus_right) lemma zmod_zminus1_eq_if: \(- a) mod b = (if a mod b = 0 then 0 else b - (a mod b))\ for a b :: int by (cases \b = 0\) (auto simp flip: minus_div_mult_eq_mod simp add: zdiv_zminus1_eq_if algebra_simps) lemma zmod_zminus2_eq_if: \a mod (- b) = (if a mod b = 0 then 0 else (a mod b) - b)\ for a b :: int by (auto simp add: zmod_zminus1_eq_if mod_minus_right) subsubsection \Borders\ lemma pos_mod_bound [simp]: "k mod l < l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_pos) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_neg) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: sgn_1_pos) qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp moreover have \int (m mod n) \ int n\ using \Suc q = n\ by simp then have \sgn s * int (m mod n) \ int n\ by (cases s \0::int\ rule: linorder_cases) simp_all ultimately show ?thesis by (simp only: modulo_int_unfold) auto qed subsubsection \Splitting Rules for div and mod\ lemma split_zdiv: \P (n div k) \ (k = 0 \ P 0) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P i)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P i))\ (is ?div) and split_zmod: \Q (n mod k) \ (k = 0 \ Q n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ Q j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ Q j))\ (is ?mod) for n k :: int proof - have *: \R (n div k) (n mod k) \ (k = 0 \ R 0 n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ R i j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ R i j))\ for R by (cases \k = 0\) (auto simp add: linorder_class.neq_iff) from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed - + text \Enable (lin)arith to deal with \<^const>\divide\ and \<^const>\modulo\ when these are applied to some constant that is of the form \<^term>\numeral k\:\ declare split_zdiv [of _ _ \numeral n\, linarith_split] for n declare split_zdiv [of _ _ \- numeral n\, linarith_split] for n declare split_zmod [of _ _ \numeral n\, linarith_split] for n declare split_zmod [of _ _ \- numeral n\, linarith_split] for n lemma zdiv_eq_0_iff: "i div k = 0 \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" (is "?L = ?R") for i k :: int proof assume ?L moreover have "?L \ ?R" by (rule split_zdiv [THEN iffD2]) simp ultimately show ?R by blast next assume ?R then show ?L by auto qed lemma zmod_trivial_iff: fixes i k :: int shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" proof - have "i mod k = i \ i div k = 0" using div_mult_mod_eq [of i k] by safe auto with zdiv_eq_0_iff show ?thesis by simp qed subsubsection \Algebraic rewrites\ lemma zdiv_zmult2_eq: \a div (b * c) = (a div b) div c\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using div_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using div_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma zdiv_zmult2_eq': \k div (l * j) = ((sgn j * k) div l) div \j\\ for k l j :: int proof - have \k div (l * j) = (sgn j * k) div (sgn j * (l * j))\ by (simp add: sgn_0_0) also have \sgn j * (l * j) = l * \j\\ by (simp add: mult.left_commute [of _ l] abs_sgn) (simp add: ac_simps) also have \(sgn j * k) div (l * \j\) = ((sgn j * k) div l) div \j\\ by (simp add: zdiv_zmult2_eq) finally show ?thesis . qed lemma zmod_zmult2_eq: \a mod (b * c) = b * (a div b mod c) + a mod b\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using mod_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using mod_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma half_nonnegative_int_iff [simp]: \k div 2 \ 0 \ k \ 0\ for k :: int by auto lemma half_negative_int_iff [simp]: \k div 2 < 0 \ k < 0\ for k :: int by auto subsubsection \Distributive laws for conversions.\ lemma zdiv_int: "int (a div b) = int a div int b" by (fact of_nat_div) lemma zmod_int: "int (a mod b) = int a mod int b" by (fact of_nat_mod) lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) +subsubsection \Monotonicity in the First Argument (Dividend)\ + +lemma zdiv_mono1: + \a div b \ a' div b\ + if \a \ a'\ \0 < b\ + for a b b' :: int +proof - + from \a \ a'\ have \b * (a div b) + a mod b \ b * (a' div b) + a' mod b\ + by simp + then have \b * (a div b) \ (a' mod b - a mod b) + b * (a' div b)\ + by (simp add: algebra_simps) + moreover have \a' mod b < b + a mod b\ + by (rule less_le_trans [of _ b]) (use \0 < b\ in simp_all) + ultimately have \b * (a div b) < b * (1 + a' div b)\ + by (simp add: distrib_left) + with \0 < b\ have \a div b < 1 + a' div b\ + by (simp add: mult_less_cancel_left) + then show ?thesis + by simp +qed + +lemma zdiv_mono1_neg: + \a' div b \ a div b\ + if \a \ a'\ \b < 0\ + for a a' b :: int + using that zdiv_mono1 [of \- a'\ \- a\ \- b\] by simp + + +subsubsection \Monotonicity in the Second Argument (Divisor)\ + +lemma zdiv_mono2: + \a div b \ a div b'\ if \0 \ a\ \0 < b'\ \b' \ b\ for a b b' :: int +proof - + define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ + then have *: \b * q + r = b' * q' + r'\ \0 \ b' * q' + r'\ + \r' < b'\ \0 \ r\ \0 < b'\ \b' \ b\ + using that by simp_all + have \0 < b' * (q' + 1)\ + using * by (simp add: distrib_left) + with * have \0 \ q'\ + by (simp add: zero_less_mult_iff) + moreover have \b * q = r' - r + b' * q'\ + using * by linarith + ultimately have \b * q < b * (q' + 1)\ + using mult_right_mono * unfolding distrib_left by fastforce + with * have \q \ q'\ + by (simp add: mult_less_cancel_left_pos) + with ** show ?thesis + by simp +qed + +lemma zdiv_mono2_neg: + \a div b' \ a div b\ if \a < 0\ \0 < b'\ \b' \ b\ for a b b' :: int +proof - + define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ + then have *: \b * q + r = b' * q' + r'\ \b' * q' + r' < 0\ + \r < b\ \0 \ r'\ \0 < b'\ \b' \ b\ + using that by simp_all + have \b' * q' < 0\ + using * by linarith + with * have \q' \ 0\ + by (simp add: mult_less_0_iff) + have \b * q' \ b' * q'\ + by (simp add: \q' \ 0\ * mult_right_mono_neg) + then have "b * q' < b * (q + 1)" + using * by (simp add: distrib_left) + then have \q' \ q\ + using * by (simp add: mult_less_cancel_left) + then show ?thesis + by (simp add: **) +qed + + +subsubsection \Quotients of Signs\ + +lemma div_eq_minus1: + \0 < b \ - 1 div b = - 1\ for b :: int + by (simp add: divide_int_def) + +lemma zmod_minus1: + \0 < b \ - 1 mod b = b - 1\ for b :: int + by (auto simp add: modulo_int_def) + +lemma minus_mod_int_eq: + \- k mod l = l - 1 - (k - 1) mod l\ if \l \ 0\ for k l :: int +proof (cases \l = 0\) + case True + then show ?thesis + by simp +next + case False + with that have \l > 0\ + by simp + then show ?thesis + proof (cases \l dvd k\) + case True + then obtain j where \k = l * j\ .. + moreover have \(l * j mod l - 1) mod l = l - 1\ + using \l > 0\ by (simp add: zmod_minus1) + then have \(l * j - 1) mod l = l - 1\ + by (simp only: mod_simps) + ultimately show ?thesis + by simp + next + case False + moreover have 1: \0 < k mod l\ + using \0 < l\ False le_less by fastforce + moreover have 2: \k mod l < 1 + l\ + using \0 < l\ pos_mod_bound[of l k] by linarith + from 1 2 \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ + by (simp add: zmod_trivial_iff) + ultimately show ?thesis + by (simp only: zmod_zminus1_eq_if) + (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) + qed +qed + +lemma div_neg_pos_less0: + \a div b < 0\ if \a < 0\ \0 < b\ for a b :: int +proof - + have "a div b \ - 1 div b" + using zdiv_mono1 that by auto + also have "... \ -1" + by (simp add: that(2) div_eq_minus1) + finally show ?thesis + by force +qed + +lemma div_nonneg_neg_le0: + \a div b \ 0\ if \0 \ a\ \b < 0\ for a b :: int + using that by (auto dest: zdiv_mono1_neg) + +lemma div_nonpos_pos_le0: + \a div b \ 0\ if \a \ 0\ \0 < b\ for a b :: int + using that by (auto dest: zdiv_mono1) + +text\Now for some equivalences of the form \a div b >=< 0 \ \\ +conditional upon the sign of \a\ or \b\. There are many more. +They should all be simp rules unless that causes too much search.\ + +lemma pos_imp_zdiv_nonneg_iff: + \0 \ a div b \ 0 \ a\ + if \0 < b\ for a b :: int +proof + assume \0 \ a div b\ + show \0 \ a\ + proof (rule ccontr) + assume \\ 0 \ a\ + then have \a < 0\ + by simp + then have \a div b < 0\ + using that by (rule div_neg_pos_less0) + with \0 \ a div b\ show False + by simp + qed +next + assume "0 \ a" + then have "0 div b \ a div b" + using zdiv_mono1 that by blast + then show "0 \ a div b" + by auto +qed + +lemma neg_imp_zdiv_nonneg_iff: + \0 \ a div b \ a \ 0\ if \b < 0\ for a b :: int + using that pos_imp_zdiv_nonneg_iff [of \- b\ \- a\] by simp + +lemma pos_imp_zdiv_pos_iff: + \0 < (i::int) div k \ k \ i\ if \0 < k\ for i k :: int + using that pos_imp_zdiv_nonneg_iff [of k i] zdiv_eq_0_iff [of i k] by arith + +lemma pos_imp_zdiv_neg_iff: + \a div b < 0 \ a < 0\ if \0 < b\ for a b :: int + \ \But not \<^prop>\a div b \ 0 \ a \ 0\; consider \<^prop>\a = 1\, \<^prop>\b = 2\ when \<^prop>\a div b = 0\.\ + using that by (simp add: pos_imp_zdiv_nonneg_iff flip: linorder_not_le) + +lemma neg_imp_zdiv_neg_iff: + \ \But not \<^prop>\a div b \ 0 \ 0 \ a\; consider \<^prop>\a = - 1\, \<^prop>\b = - 2\ when \<^prop>\a div b = 0\.\ + \a div b < 0 \ 0 < a\ if \b < 0\ for a b :: int + using that by (simp add: neg_imp_zdiv_nonneg_iff flip: linorder_not_le) + +lemma nonneg1_imp_zdiv_pos_iff: + \a div b > 0 \ a \ b \ b > 0\ if \0 \ a\ for a b :: int +proof - + have "0 < a div b \ b \ a" + using div_pos_pos_trivial[of a b] that by arith + moreover have "0 < a div b \ b > 0" + using that div_nonneg_neg_le0[of a b] by (cases "b=0"; force) + moreover have "b \ a \ 0 < b \ 0 < a div b" + using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp + ultimately show ?thesis + by blast +qed + +lemma zmod_le_nonneg_dividend: + \m mod k \ m\ if \(m::int) \ 0\ for m k :: int +proof - + from that have \m > 0 \ m = 0\ + by auto + then show ?thesis proof + assume \m = 0\ then show ?thesis + by simp + next + assume \m > 0\ then show ?thesis + proof (cases k \0::int\ rule: linorder_cases) + case less + moreover define l where \l = - k\ + ultimately have \l > 0\ + by simp + with \m > 0\ have \int (nat m mod nat l) \ m\ + by (simp flip: le_nat_iff) + then have \int (nat m mod nat l) - l \ m\ + using \l > 0\ by simp + with \m > 0\ \l > 0\ show ?thesis + by (simp add: modulo_int_def l_def flip: le_nat_iff) + qed (simp_all add: modulo_int_def flip: le_nat_iff) + qed +qed + +lemma sgn_div_eq_sgn_mult: + \sgn (k div l) = of_bool (k div l \ 0) * sgn (k * l)\ + for k l :: int +proof (cases \k div l = 0\) + case True + then show ?thesis + by simp +next + case False + have \0 \ \k\ div \l\\ + by (cases \l = 0\) (simp_all add: pos_imp_zdiv_nonneg_iff) + then have \\k\ div \l\ \ 0 \ 0 < \k\ div \l\\ + by (simp add: less_le) + also have \\ \ \k\ \ \l\\ + using False nonneg1_imp_zdiv_pos_iff by auto + finally have *: \\k\ div \l\ \ 0 \ \l\ \ \k\\ . + show ?thesis + using \0 \ \k\ div \l\\ False + by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l] + sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp) +qed + + +subsubsection \Further properties\ + +lemma div_int_pos_iff: + "k div l \ 0 \ k = 0 \ l = 0 \ k \ 0 \ l \ 0 + \ k < 0 \ l < 0" + for k l :: int +proof (cases "k = 0 \ l = 0") + case False + then have *: "k \ 0" "l \ 0" + by auto + then have "0 \ k div l \ \ k < 0 \ 0 \ l" + by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) + then show ?thesis + using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) +qed auto + +lemma mod_int_pos_iff: + \k mod l \ 0 \ l dvd k \ l = 0 \ k \ 0 \ l > 0\ + for k l :: int +proof (cases "l > 0") + case False + then show ?thesis + by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \auto simp add: le_less not_less\) +qed auto + +lemma abs_div: + \\x div y\ = \x\ div \y\\ if \y dvd x\ for x y :: int + using that by (cases \y = 0\) (auto simp add: abs_mult) + +lemma int_power_div_base: \<^marker>\contributor \Matthias Daum\\ + \k ^ m div k = k ^ (m - Suc 0)\ if \0 < m\ \0 < k\ for k :: int + using that by (cases m) simp_all + +lemma int_div_less_self: \<^marker>\contributor \Matthias Daum\\ + \x div k < x\ if \0 < x\ \1 < k\ for x k :: int +proof - + from that have \nat (x div k) = nat x div nat k\ + by (simp add: nat_div_distrib) + also from that have \nat x div nat k < nat x\ + by simp + finally show ?thesis + by simp +qed + + +subsubsection \Computing \div\ and \mod\ by shifting\ + +lemma div_pos_geq: + \k div l = (k - l) div l + 1\ if \0 < l\ \l \ k\ for k l :: int +proof - + have "k = (k - l) + l" by simp + then obtain j where k: "k = j + l" .. + with that show ?thesis by (simp add: div_add_self2) +qed + +lemma mod_pos_geq: + \k mod l = (k - l) mod l\ if \0 < l\ \l \ k\ for k l :: int +proof - + have "k = (k - l) + l" by simp + then obtain j where k: "k = j + l" .. + with that show ?thesis by simp +qed + +lemma pos_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = b div a\ (is ?Q) + and pos_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\ (is ?R) + if \0 \ a\ for a b :: int +proof - + have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\ + proof (cases \2 * a\ \b div a\ \1 + 2 * (b mod a)\ \1 + 2 * b\ rule: euclidean_relation_intI) + case by0 + then show ?case + by simp + next + case divides + have \2 dvd (2 * a)\ + by simp + then have \2 dvd (1 + 2 * b)\ + using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) + then have \2 dvd (1 + b * 2)\ + by (simp add: ac_simps) + then have \is_unit (2 :: int)\ + by simp + then show ?case + by simp + next + case euclidean_relation + with that have \a > 0\ + by simp + moreover have \b mod a < a\ + using \a > 0\ by simp + then have \1 + 2 * (b mod a) < 2 * a\ + by simp + moreover have \2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\ + by (simp only: algebra_simps) + moreover have \0 \ 2 * (b mod a)\ + using \a > 0\ by simp + ultimately show ?case + by (simp add: algebra_simps) + qed + then show ?Q and ?R + by simp_all +qed + +lemma neg_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = (b + 1) div a\ (is ?Q) + and neg_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\ (is ?R) + if \a \ 0\ for a b :: int +proof - + have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\ + proof (cases \2 * a\ \(b + 1) div a\ \2 * ((b + 1) mod a) - 1\ \1 + 2 * b\ rule: euclidean_relation_intI) + case by0 + then show ?case + by simp + next + case divides + have \2 dvd (2 * a)\ + by simp + then have \2 dvd (1 + 2 * b)\ + using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) + then have \2 dvd (1 + b * 2)\ + by (simp add: ac_simps) + then have \is_unit (2 :: int)\ + by simp + then show ?case + by simp + next + case euclidean_relation + with that have \a < 0\ + by simp + moreover have \(b + 1) mod a > a\ + using \a < 0\ by simp + then have \2 * ((b + 1) mod a) > 1 + 2 * a\ + by simp + moreover have \((1 + b) mod a) \ 0\ + using \a < 0\ by simp + then have \2 * ((1 + b) mod a) \ 0\ + by simp + moreover have \2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) = + 2 * ((1 + b) div a * a + (1 + b) mod a)\ + by (simp only: algebra_simps) + ultimately show ?case + by (simp add: algebra_simps sgn_mult abs_mult) + qed + then show ?Q and ?R + by simp_all +qed + +lemma zdiv_numeral_Bit0 [simp]: + \numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = + numeral v div (numeral w :: int)\ + unfolding numeral.simps unfolding mult_2 [symmetric] + by (rule div_mult_mult1) simp + +lemma zdiv_numeral_Bit1 [simp]: + \numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = + (numeral v div (numeral w :: int))\ + unfolding numeral.simps + unfolding mult_2 [symmetric] add.commute [of _ 1] + by (rule pos_zdiv_mult_2) simp + +lemma zmod_numeral_Bit0 [simp]: + \numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = + (2::int) * (numeral v mod numeral w)\ + unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] + unfolding mult_2 [symmetric] by (rule mod_mult_mult1) + +lemma zmod_numeral_Bit1 [simp]: + \numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = + 2 * (numeral v mod numeral w) + (1::int)\ + unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] + unfolding mult_2 [symmetric] add.commute [of _ 1] + by (rule pos_zmod_mult_2) simp + + subsection \Generic symbolic computations\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. \ class unique_euclidean_semiring_with_nat_division = unique_euclidean_semiring_with_nat + fixes divmod :: \num \ num \ 'a \ 'a\ and divmod_step :: \'a \ 'a \ 'a \ 'a \ 'a\ \ \ These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ assumes divmod_def: \divmod m n = (numeral m div numeral n, numeral m mod numeral n)\ and divmod_step_def [simp]: \divmod_step l (q, r) = (if euclidean_size l \ euclidean_size r then (2 * q + 1, r - l) else (2 * q, r))\ \ \ This is a formulation of one step (referring to one digit position) in school-method division: compare the dividend at the current digit position with the remainder from previous division steps and evaluate accordingly.\ begin lemma fst_divmod: \fst (divmod m n) = numeral m div numeral n\ by (simp add: divmod_def) lemma snd_divmod: \snd (divmod m n) = numeral m mod numeral n\ by (simp add: divmod_def) text \ Following a formulation of school-method division. If the divisor is smaller than the dividend, terminate. If not, shift the dividend to the right until termination occurs and then reiterate single division steps in the opposite direction. \ lemma divmod_divmod_step: \divmod m n = (if m < n then (0, numeral m) else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\ proof (cases \m < n\) case True then show ?thesis by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod) next case False define r s t where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ \t = 2 * s\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 n) = of_nat t\ and \\ s \ r mod s\ by (simp_all add: not_le) have t: \2 * (r div t) = r div s - r div s mod 2\ \r mod t = s * (r div s mod 2) + r mod s\ by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Division.div_mult2_eq \t = 2 * s\) (use mod_mult2_eq [of r s 2] in \simp add: ac_simps \t = 2 * s\\) have rs: \r div s mod 2 = 0 \ r div s mod 2 = Suc 0\ by auto from \\ s \ r mod s\ have \s \ r mod t \ r div s = Suc (2 * (r div t)) \ r mod s = r mod t - s\ using rs by (auto simp add: t) moreover have \r mod t < s \ r div s = 2 * (r div t) \ r mod s = r mod t\ using rs by (auto simp add: t) ultimately show ?thesis by (simp add: divmod_def prod_eq_iff split_def Let_def not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *) (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff) qed text \The division rewrite proper -- first, trivial results involving \1\\ lemma divmod_trivial [simp]: "divmod m Num.One = (numeral m, 0)" "divmod num.One (num.Bit0 n) = (0, Numeral1)" "divmod num.One (num.Bit1 n) = (0, Numeral1)" using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) text \Division by an even number is a right-shift\ lemma divmod_cancel [simp]: \divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r))\ (is ?P) \divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r + 1))\ (is ?Q) proof - define r s where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 m) = of_nat (2 * r)\ \numeral (num.Bit0 n) = of_nat (2 * s)\ \numeral (num.Bit1 m) = of_nat (Suc (2 * r))\ by simp_all have **: \Suc (2 * r) div 2 = r\ by simp show ?P and ?Q by (simp_all add: divmod_def *) (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc add: Euclidean_Division.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **) qed text \The really hard work\ lemma divmod_steps [simp]: "divmod (num.Bit0 m) (num.Bit1 n) = (if m \ n then (0, numeral (num.Bit0 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit0 m) (num.Bit0 (num.Bit1 n))))" "divmod (num.Bit1 m) (num.Bit1 n) = (if m < n then (0, numeral (num.Bit1 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps text \Special case: divisibility\ definition divides_aux :: "'a \ 'a \ bool" where "divides_aux qr \ snd qr = 0" lemma divides_aux_eq [simp]: "divides_aux (q, r) \ r = 0" by (simp add: divides_aux_def) lemma dvd_numeral_simp [simp]: "numeral m dvd numeral n \ divides_aux (divmod n m)" by (simp add: divmod_def mod_eq_0_iff_dvd) -text \Generic computation of quotient and remainder\ - -lemma numeral_div_numeral [simp]: +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]: +lemma numeral_mod_numeral [simp]: "numeral k mod numeral l = snd (divmod k l)" by (simp add: snd_divmod) lemma one_div_numeral [simp]: "1 div numeral n = fst (divmod num.One n)" by (simp add: fst_divmod) lemma one_mod_numeral [simp]: "1 mod numeral n = snd (divmod num.One n)" by (simp add: snd_divmod) end instantiation nat :: unique_euclidean_semiring_with_nat_division begin definition divmod_nat :: "num \ num \ nat \ nat" where divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_nat :: "nat \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ l then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (simp_all add: divmod'_nat_def divmod_step_nat_def) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: \Suc 0 div numeral Num.One = 1\ \Suc 0 div numeral (Num.Bit0 n) = 0\ \Suc 0 div numeral (Num.Bit1 n) = 0\ by simp_all lemma Suc_0_mod_numeral [simp]: \Suc 0 mod numeral Num.One = 0\ \Suc 0 mod numeral (Num.Bit0 n) = 1\ \Suc 0 mod numeral (Num.Bit1 n) = 1\ by simp_all instantiation int :: unique_euclidean_semiring_with_nat_division begin definition divmod_int :: "num \ num \ int \ int" where "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_int :: "int \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if \l\ \ \r\ then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (auto simp add: divmod_int_def divmod_step_int_def) end declare divmod_algorithm_code [where ?'a = int, code] context begin - + qualified definition adjust_div :: "int \ int \ int" where "adjust_div qr = (let (q, r) = qr in q + of_bool (r \ 0))" qualified lemma adjust_div_eq [simp, code]: "adjust_div (q, r) = q + of_bool (r \ 0)" by (simp add: adjust_div_def) qualified definition adjust_mod :: "num \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)" lemma minus_numeral_div_numeral [simp]: "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma minus_numeral_mod_numeral [simp]: "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma numeral_div_minus_numeral [simp]: "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed - + lemma numeral_mod_minus_numeral [simp]: "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma minus_one_div_numeral [simp]: "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" - using minus_numeral_div_numeral [of Num.One n] by simp + using minus_numeral_div_numeral [of Num.One n] by simp lemma minus_one_mod_numeral [simp]: "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)" using minus_numeral_mod_numeral [of Num.One n] by simp lemma one_div_minus_numeral [simp]: "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" using numeral_div_minus_numeral [of Num.One n] by simp - + lemma one_mod_minus_numeral [simp]: "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp lemma [code]: fixes k :: int - shows + shows "k div 0 = 0" "k mod 0 = k" "0 div k = 0" "0 mod k = 0" "k div Int.Pos Num.One = k" "k mod Int.Pos Num.One = 0" "k div Int.Neg Num.One = - k" "k mod Int.Neg Num.One = 0" "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)" "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" by simp_all end lemma divmod_BitM_2_eq [simp]: \divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\ by (cases m) simp_all subsubsection \Computation by simplification\ lemma euclidean_size_nat_less_eq_iff: \euclidean_size m \ euclidean_size n \ m \ n\ for m n :: nat by simp lemma euclidean_size_int_less_eq_iff: \euclidean_size k \ euclidean_size l \ \k\ \ \l\\ for k l :: int by auto simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | "numeral a div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \ let val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\If\); fun successful_rewrite ctxt ct = let val thm = Simplifier.rewrite ctxt ct in if Thm.is_reflexive thm then NONE else SOME thm end; in fn phi => let val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral one_div_minus_numeral one_mod_minus_numeral numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral numeral_div_minus_numeral numeral_mod_minus_numeral div_minus_minus mod_minus_minus Euclidean_Division.adjust_div_eq of_bool_eq one_neq_zero numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One case_prod_beta rel_simps Euclidean_Division.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral} @ [@{lemma "0 = 0 \ True" by simp}]); fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end end \ \ \ There is space for improvement here: the calculation itself could be carried out outside the logic, and a generic simproc - (simplifier setup) for generic calculation would be helpful. + (simplifier setup) for generic calculation would be helpful. \ subsubsection \Code generation\ context begin qualified definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" qualified lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = divmod_nat (m - n) n in (Suc q, r))" by (simp add: divmod_nat_def prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) qualified lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by (simp_all add: divmod_nat_def) end code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Groebner_Basis.thy b/src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy +++ b/src/HOL/Groebner_Basis.thy @@ -1,96 +1,100 @@ (* Title: HOL/Groebner_Basis.thy Author: Amine Chaieb, TU Muenchen *) section \Groebner bases\ theory Groebner_Basis imports Semiring_Normalization Parity begin subsection \Groebner Bases\ lemmas bool_simps = simp_thms(1-34) \ \FIXME move to \<^theory>\HOL.HOL\\ lemma nnf_simps: \ \FIXME shadows fact binding in \<^theory>\HOL.HOL\\ "(\(P \ Q)) = (\P \ \Q)" "(\(P \ Q)) = (\P \ \Q)" "(P \ Q) = (\P \ Q)" "(P = Q) = ((P \ Q) \ (\P \ \ Q))" "(\ \(P)) = P" by blast+ lemma dnf: "(P \ (Q \ R)) = ((P\Q) \ (P\R))" "((Q \ R) \ P) = ((Q\P) \ (R\P))" "(P \ Q) = (Q \ P)" "(P \ Q) = (Q \ P)" by blast+ lemmas weak_dnf_simps = dnf bool_simps lemma PFalse: "P \ False \ \ P" "\ P \ (P \ False)" by auto named_theorems algebra "pre-simplification rules for algebraic methods" ML_file \Tools/groebner.ML\ method_setup algebra = \ let fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () val addN = "add" val delN = "del" val any_keyword = keyword addN || keyword delN val thms = Scan.repeats (Scan.unless any_keyword Attrib.multi_thm); in Scan.optional (keyword addN |-- thms) [] -- Scan.optional (keyword delN |-- thms) [] >> (fn (add_ths, del_ths) => fn ctxt => SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt)) end \ "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases" declare dvd_def[algebra] declare mod_eq_0_iff_dvd[algebra] declare mod_div_trivial[algebra] declare mod_mod_trivial[algebra] declare div_by_0[algebra] declare mod_by_0[algebra] declare mult_div_mod_eq[algebra] declare div_minus_minus[algebra] declare mod_minus_minus[algebra] declare div_minus_right[algebra] declare mod_minus_right[algebra] declare div_0[algebra] declare mod_0[algebra] declare mod_by_1[algebra] declare div_by_1[algebra] declare mod_minus1_right[algebra] declare div_minus1_right[algebra] declare mod_mult_self2_is_0[algebra] declare mod_mult_self1_is_0[algebra] -declare zmod_eq_0_iff[algebra] + +lemma zmod_eq_0_iff [algebra]: + \m mod d = 0 \ (\q. m = d * q)\ for m d :: int + by (auto simp add: mod_eq_0_iff_dvd) + declare dvd_0_left_iff[algebra] declare zdvd1_eq[algebra] declare mod_eq_dvd_iff[algebra] declare nat_mod_eq_iff[algebra] context semiring_parity begin declare even_mult_iff [algebra] declare even_power [algebra] end context ring_parity begin declare even_minus [algebra] end declare even_Suc [algebra] declare even_diff_nat [algebra] end diff --git a/src/HOL/Main.thy b/src/HOL/Main.thy --- a/src/HOL/Main.thy +++ b/src/HOL/Main.thy @@ -1,108 +1,109 @@ section \Main HOL\ text \ Classical Higher-order Logic -- only ``Main'', excluding real and complex numbers etc. \ theory Main imports Predicate_Compile Quickcheck_Narrowing Mirabelle Extraction Nunchaku BNF_Greatest_Fixpoint Filter Conditionally_Complete_Lattices Binomial GCD + Divides begin subsection \Namespace cleanup\ hide_const (open) czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift shift proj id_bnf hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV subsection \Syntax cleanup\ no_notation ordLeq2 (infix "<=o" 50) and ordLeq3 (infix "\o" 50) and ordLess2 (infix "(_,/ _)\") bundle cardinal_syntax begin notation ordLeq2 (infix "<=o" 50) and ordLeq3 (infix "\o" 50) and ordLess2 (infix "Lattice syntax\ bundle lattice_syntax begin notation bot ("\") and top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\ _" [900] 900) and Sup ("\ _" [900] 900) syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) end bundle no_lattice_syntax begin no_notation bot ("\") and top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\ _" [900] 900) and Sup ("\ _" [900] 900) no_syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) end unbundle no_lattice_syntax end diff --git a/src/HOL/Number_Theory/Cong.thy b/src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy +++ b/src/HOL/Number_Theory/Cong.thy @@ -1,770 +1,754 @@ (* Title: HOL/Number_Theory/Cong.thy Author: Christophe Tabacznyj Author: Lawrence C. Paulson Author: Amine Chaieb Author: Thomas M. Rasmussen Author: Jeremy Avigad Defines congruence (notation: [x = y] (mod z)) for natural numbers and integers. This file combines and revises a number of prior developments. The original theories "GCD" and "Primes" were by Christophe Tabacznyj and Lawrence C. Paulson, based on @{cite davenport92}. They introduced gcd, lcm, and prime for the natural numbers. The original theory "IntPrimes" was by Thomas M. Rasmussen, and extended gcd, lcm, primes to the integers. Amine Chaieb provided another extension of the notions to the integers, and added a number of results to "Primes" and "GCD". The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and developed the congruence relations on the integers. The notion was extended to the natural numbers by Chaieb. Jeremy Avigad combined these, revised and tidied them, made the development uniform for the natural numbers and the integers, and added a number of new theorems. *) section \Congruence\ theory Cong imports "HOL-Computational_Algebra.Primes" begin subsection \Generic congruences\ context unique_euclidean_semiring begin definition cong :: "'a \ 'a \ 'a \ bool" (\(1[_ = _] '(' mod _'))\) where "cong b c a \ b mod a = c mod a" abbreviation notcong :: "'a \ 'a \ 'a \ bool" (\(1[_ \ _] '(' mod _'))\) where "notcong b c a \ \ cong b c a" lemma cong_refl [simp]: "[b = b] (mod a)" by (simp add: cong_def) lemma cong_sym: "[b = c] (mod a) \ [c = b] (mod a)" by (simp add: cong_def) lemma cong_sym_eq: "[b = c] (mod a) \ [c = b] (mod a)" by (auto simp add: cong_def) lemma cong_trans [trans]: "[b = c] (mod a) \ [c = d] (mod a) \ [b = d] (mod a)" by (simp add: cong_def) lemma cong_mult_self_right: "[b * a = 0] (mod a)" by (simp add: cong_def) lemma cong_mult_self_left: "[a * b = 0] (mod a)" by (simp add: cong_def) lemma cong_mod_left [simp]: "[b mod a = c] (mod a) \ [b = c] (mod a)" by (simp add: cong_def) lemma cong_mod_right [simp]: "[b = c mod a] (mod a) \ [b = c] (mod a)" by (simp add: cong_def) lemma cong_0 [simp, presburger]: "[b = c] (mod 0) \ b = c" by (simp add: cong_def) lemma cong_1 [simp, presburger]: "[b = c] (mod 1)" by (simp add: cong_def) lemma cong_dvd_iff: "a dvd b \ a dvd c" if "[b = c] (mod a)" using that by (auto simp: cong_def dvd_eq_mod_eq_0) lemma cong_0_iff: "[b = 0] (mod a) \ a dvd b" by (simp add: cong_def dvd_eq_mod_eq_0) lemma cong_add: "[b = c] (mod a) \ [d = e] (mod a) \ [b + d = c + e] (mod a)" by (auto simp add: cong_def intro: mod_add_cong) lemma cong_mult: "[b = c] (mod a) \ [d = e] (mod a) \ [b * d = c * e] (mod a)" by (auto simp add: cong_def intro: mod_mult_cong) lemma cong_scalar_right: "[b = c] (mod a) \ [b * d = c * d] (mod a)" by (simp add: cong_mult) lemma cong_scalar_left: "[b = c] (mod a) \ [d * b = d * c] (mod a)" by (simp add: cong_mult) lemma cong_pow: "[b = c] (mod a) \ [b ^ n = c ^ n] (mod a)" by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a]) lemma cong_sum: "[sum f A = sum g A] (mod a)" if "\x. x \ A \ [f x = g x] (mod a)" using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add) lemma cong_prod: "[prod f A = prod g A] (mod a)" if "(\x. x \ A \ [f x = g x] (mod a))" using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult) lemma mod_mult_cong_right: "[c mod (a * b) = d] (mod a) \ [c = d] (mod a)" by (simp add: cong_def mod_mod_cancel mod_add_left_eq) lemma mod_mult_cong_left: "[c mod (b * a) = d] (mod a) \ [c = d] (mod a)" using mod_mult_cong_right [of c a b d] by (simp add: ac_simps) end context unique_euclidean_ring begin lemma cong_diff: "[b = c] (mod a) \ [d = e] (mod a) \ [b - d = c - e] (mod a)" by (auto simp add: cong_def intro: mod_diff_cong) lemma cong_diff_iff_cong_0: "[b - c = 0] (mod a) \ [b = c] (mod a)" (is "?P \ ?Q") proof assume ?P then have "[b - c + c = 0 + c] (mod a)" by (rule cong_add) simp then show ?Q by simp next assume ?Q with cong_diff [of b c a c c] show ?P by simp qed lemma cong_minus_minus_iff: "[- b = - c] (mod a) \ [b = c] (mod a)" using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a] by (simp add: cong_0_iff dvd_diff_commute) lemma cong_modulus_minus_iff [iff]: "[b = c] (mod - a) \ [b = c] (mod a)" using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"] by (simp add: cong_0_iff) lemma cong_iff_dvd_diff: "[a = b] (mod m) \ m dvd (a - b)" by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0) lemma cong_iff_lin: "[a = b] (mod m) \ (\k. b = a + m * k)" (is "?P \ ?Q") proof - have "?P \ m dvd b - a" by (simp add: cong_iff_dvd_diff dvd_diff_commute) also have "\ \ ?Q" by (auto simp add: algebra_simps elim!: dvdE) finally show ?thesis by simp qed lemma cong_add_lcancel: "[a + x = a + y] (mod n) \ [x = y] (mod n)" by (simp add: cong_iff_lin algebra_simps) lemma cong_add_rcancel: "[x + a = y + a] (mod n) \ [x = y] (mod n)" by (simp add: cong_iff_lin algebra_simps) lemma cong_add_lcancel_0: "[a + x = a] (mod n) \ [x = 0] (mod n)" using cong_add_lcancel [of a x 0 n] by simp lemma cong_add_rcancel_0: "[x + a = a] (mod n) \ [x = 0] (mod n)" using cong_add_rcancel [of x a 0 n] by simp lemma cong_dvd_modulus: "[x = y] (mod n)" if "[x = y] (mod m)" and "n dvd m" using that by (auto intro: dvd_trans simp add: cong_iff_dvd_diff) lemma cong_modulus_mult: "[x = y] (mod m)" if "[x = y] (mod m * n)" using that by (simp add: cong_iff_dvd_diff) (rule dvd_mult_left) end lemma cong_abs [simp]: "[x = y] (mod \m\) \ [x = y] (mod m)" for x y :: "'a :: {unique_euclidean_ring, linordered_idom}" by (simp add: cong_iff_dvd_diff) lemma cong_square: "prime p \ 0 < a \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = - 1] (mod p)" for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}" by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest: prime_dvd_multD) lemma cong_mult_rcancel: "[a * k = b * k] (mod m) \ [a = b] (mod m)" if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" using that by (auto simp add: cong_iff_dvd_diff left_diff_distrib [symmetric] ac_simps coprime_dvd_mult_right_iff) lemma cong_mult_lcancel: "[k * a = k * b] (mod m) = [a = b] (mod m)" if "coprime k m" for a k m :: "'a::{unique_euclidean_ring, ring_gcd}" using that cong_mult_rcancel [of k m a b] by (simp add: ac_simps) lemma coprime_cong_mult: "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" for a b :: "'a :: {unique_euclidean_ring, semiring_gcd}" by (simp add: cong_iff_dvd_diff divides_mult) lemma cong_gcd_eq: "gcd a m = gcd b m" if "[a = b] (mod m)" for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" proof (cases "m = 0") case True with that show ?thesis by simp next case False moreover have "gcd (a mod m) m = gcd (b mod m) m" using that by (simp add: cong_def) ultimately show ?thesis by simp qed lemma cong_imp_coprime: "[a = b] (mod m) \ coprime a m \ coprime b m" for a b :: "'a :: {unique_euclidean_semiring, euclidean_semiring_gcd}" by (auto simp add: coprime_iff_gcd_eq_1 dest: cong_gcd_eq) lemma cong_cong_prod_coprime: "[x = y] (mod (\i\A. m i))" if "(\i\A. [x = y] (mod m i))" "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" for x y :: "'a :: {unique_euclidean_ring, semiring_gcd}" using that by (induct A rule: infinite_finite_induct) (auto intro!: coprime_cong_mult prod_coprime_right) subsection \Congruences on \<^typ>\nat\ and \<^typ>\int\\ lemma cong_int_iff: "[int m = int q] (mod int n) \ [m = q] (mod n)" by (simp add: cong_def of_nat_mod [symmetric]) lemma cong_Suc_0 [simp, presburger]: "[m = n] (mod Suc 0)" using cong_1 [of m n] by simp lemma cong_diff_nat: "[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)" and "a \ c" "b \ d" for a b c d m :: nat proof - have "[c + (a - c) = d + (b - d)] (mod m)" using that by simp with \[c = d] (mod m)\ have "[c + (a - c) = c + (b - d)] (mod m)" using mod_add_cong by (auto simp add: cong_def) fastforce then show ?thesis by (simp add: cong_def nat_mod_eq_iff) qed lemma cong_diff_iff_cong_0_nat: "[a - b = 0] (mod m) \ [a = b] (mod m)" if "a \ b" for a b :: nat - using that by (auto simp add: cong_def le_imp_diff_is_add dest: nat_mod_eq_lemma) + using that by (simp add: cong_0_iff) (simp add: cong_def mod_eq_dvd_iff_nat) lemma cong_diff_iff_cong_0_nat': "[nat \int a - int b\ = 0] (mod m) \ [a = b] (mod m)" proof (cases "b \ a") case True then show ?thesis by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of b a m]) next case False then have "a \ b" by simp then show ?thesis by (simp add: nat_diff_distrib' cong_diff_iff_cong_0_nat [of a b m]) (auto simp add: cong_def) qed lemma cong_altdef_nat: "a \ b \ [a = b] (mod m) \ m dvd (a - b)" for a b :: nat by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat) lemma cong_altdef_nat': "[a = b] (mod m) \ m dvd nat \int a - int b\" using cong_diff_iff_cong_0_nat' [of a b m] by (simp only: cong_0_iff [symmetric]) lemma cong_mult_rcancel_nat: "[a * k = b * k] (mod m) \ [a = b] (mod m)" if "coprime k m" for a k m :: nat proof - have "[a * k = b * k] (mod m) \ m dvd nat \int (a * k) - int (b * k)\" by (simp add: cong_altdef_nat') also have "\ \ m dvd nat \(int a - int b) * int k\" by (simp add: algebra_simps) also have "\ \ m dvd nat \int a - int b\ * k" by (simp add: abs_mult nat_times_as_int) also have "\ \ m dvd nat \int a - int b\" by (rule coprime_dvd_mult_left_iff) (use \coprime k m\ in \simp add: ac_simps\) also have "\ \ [a = b] (mod m)" by (simp add: cong_altdef_nat') finally show ?thesis . qed lemma cong_mult_lcancel_nat: "[k * a = k * b] (mod m) = [a = b] (mod m)" if "coprime k m" for a k m :: nat using that by (simp add: cong_mult_rcancel_nat ac_simps) lemma coprime_cong_mult_nat: "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" for a b :: nat by (simp add: cong_altdef_nat' divides_mult) lemma cong_less_imp_eq_nat: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: nat by (auto simp add: cong_def) lemma cong_less_imp_eq_int: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: int by (auto simp add: cong_def) lemma cong_less_unique_nat: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: nat by (auto simp: cong_def) (metis mod_less_divisor mod_mod_trivial) lemma cong_less_unique_int: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: int by (auto simp: cong_def) (metis mod_mod_trivial pos_mod_conj) -lemma cong_iff_lin_nat: "([a = b] (mod m)) \ (\k1 k2. b + k1 * m = a + k2 * m)" - (is "?lhs = ?rhs") +lemma cong_iff_lin_nat: "[a = b] (mod m) \ (\k1 k2. b + k1 * m = a + k2 * m)" for a b :: nat -proof - assume ?lhs - show ?rhs - proof (cases "b \ a") - case True - with \?lhs\ show ?rhs - by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult.commute) - next - case False - with \?lhs\ show ?rhs - by (metis cong_def mult.commute nat_le_linear nat_mod_eq_lemma) - qed -next - assume ?rhs - then show ?lhs - by (metis cong_def mult.commute nat_mod_eq_iff) -qed + apply (auto simp add: cong_def nat_mod_eq_iff) + apply (metis mult.commute) + apply (metis mult.commute) + done lemma cong_cong_mod_nat: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: nat by simp lemma cong_cong_mod_int: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: int by simp lemma cong_add_lcancel_nat: "[a + x = a + y] (mod n) \ [x = y] (mod n)" for a x y :: nat by (simp add: cong_iff_lin_nat) lemma cong_add_rcancel_nat: "[x + a = y + a] (mod n) \ [x = y] (mod n)" for a x y :: nat by (simp add: cong_iff_lin_nat) lemma cong_add_lcancel_0_nat: "[a + x = a] (mod n) \ [x = 0] (mod n)" for a x :: nat using cong_add_lcancel_nat [of a x 0 n] by simp lemma cong_add_rcancel_0_nat: "[x + a = a] (mod n) \ [x = 0] (mod n)" for a x :: nat using cong_add_rcancel_nat [of x a 0 n] by simp lemma cong_dvd_modulus_nat: "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" for x y :: nat - unfolding cong_iff_lin_nat dvd_def - by (metis mult.commute mult.left_commute) + by (auto simp add: cong_altdef_nat') lemma cong_to_1_nat: fixes a :: nat assumes "[a = 1] (mod n)" shows "n dvd (a - 1)" proof (cases "a = 0") case True then show ?thesis by force next case False with assms show ?thesis by (metis cong_altdef_nat leI less_one) qed lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \ n = Suc 0" by (auto simp: cong_def) lemma cong_0_1_nat: "[0 = 1] (mod n) \ n = 1" for n :: nat by (auto simp: cong_def) lemma cong_0_1_int: "[0 = 1] (mod n) \ n = 1 \ n = - 1" for n :: int by (auto simp: cong_def zmult_eq_1_iff) lemma cong_to_1'_nat: "[a = 1] (mod n) \ a = 0 \ n = 1 \ (\m. a = 1 + m * n)" for a :: nat by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) lemma cong_le_nat: "y \ x \ [x = y] (mod n) \ (\q. x = q * n + y)" for x y :: nat - by (auto simp add: cong_altdef_nat le_imp_diff_is_add elim!: dvdE) - + by (auto simp add: cong_altdef_nat le_imp_diff_is_add) lemma cong_solve_nat: fixes a :: nat shows "\x. [a * x = gcd a n] (mod n)" proof (cases "a = 0 \ n = 0") case True then show ?thesis by (force simp add: cong_0_iff cong_sym) next case False then show ?thesis using bezout_nat [of a n] by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) qed lemma cong_solve_int: fixes a :: int shows "\x. [a * x = gcd a n] (mod n)" by (metis bezout_int cong_iff_lin mult.commute) lemma cong_solve_dvd_nat: fixes a :: nat assumes "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - from cong_solve_nat [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast also from assms have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" by auto finally show ?thesis by auto qed lemma cong_solve_dvd_int: fixes a::int assumes b: "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - from cong_solve_int [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast also from b have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" by auto finally show ?thesis by auto qed lemma cong_solve_coprime_nat: "\x. [a * x = Suc 0] (mod n)" if "coprime a n" using that cong_solve_nat [of a n] by auto lemma cong_solve_coprime_int: "\x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int using that cong_solve_int [of a n] by (auto simp add: zabs_def split: if_splits) lemma coprime_iff_invertible_nat: "coprime a m \ (\x. [a * x = Suc 0] (mod m))" (is "?P \ ?Q") proof assume ?P then show ?Q by (auto dest!: cong_solve_coprime_nat) next assume ?Q then obtain b where "[a * b = Suc 0] (mod m)" by blast with coprime_mod_left_iff [of m "a * b"] show ?P by (cases "m = 0 \ m = 1") (unfold cong_def, auto simp add: cong_def) qed lemma coprime_iff_invertible_int: "coprime a m \ (\x. [a * x = 1] (mod m))" (is "?P \ ?Q") for m :: int proof assume ?P then show ?Q by (auto dest: cong_solve_coprime_int) next assume ?Q then obtain b where "[a * b = 1] (mod m)" by blast with coprime_mod_left_iff [of m "a * b"] show ?P by (cases "m = 0 \ m = 1") (unfold cong_def, auto simp add: zmult_eq_1_iff) qed lemma coprime_iff_invertible'_nat: assumes "m > 0" shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = Suc 0] (mod m))" proof - have "\b. \0 < m; [a * b = Suc 0] (mod m)\ \ \b' 0" shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = 1] (mod m))" proof - have "\b. \0 < m; [a * b = 1] (mod m)\ \ \b' [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: nat by (meson cong_altdef_nat' lcm_least) lemma cong_cong_lcm_int: "[x = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: int by (auto simp add: cong_iff_dvd_diff lcm_least) lemma cong_cong_prod_coprime_nat: "[x = y] (mod (\i\A. m i))" if "(\i\A. [x = y] (mod m i))" "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" for x y :: nat using that by (induct A rule: infinite_finite_induct) (auto intro!: coprime_cong_mult_nat prod_coprime_right) lemma binary_chinese_remainder_nat: fixes m1 m2 :: nat assumes a: "coprime m1 m2" shows "\x. [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - have "\b1 b2. [b1 = 1] (mod m1) \ [b1 = 0] (mod m2) \ [b2 = 0] (mod m1) \ [b2 = 1] (mod m2)" proof - from cong_solve_coprime_nat [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" by (simp add: ac_simps) from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed then obtain b1 b2 where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis by blast qed lemma binary_chinese_remainder_int: fixes m1 m2 :: int assumes a: "coprime m1 m2" shows "\x. [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - have "\b1 b2. [b1 = 1] (mod m1) \ [b1 = 0] (mod m2) \ [b2 = 0] (mod m1) \ [b2 = 1] (mod m2)" proof - from cong_solve_coprime_int [OF a] obtain x1 where 1: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" by (simp add: ac_simps) from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed then obtain b1 b2 where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)" by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis by blast qed lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \ [x = y] (mod m)" for x y :: nat by (metis cong_def mod_mult_cong_right) lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \ x < m \ y < m \ x = y" for x y :: nat by (simp add: cong_def) lemma binary_chinese_remainder_unique_nat: fixes m1 m2 :: nat assumes a: "coprime m1 m2" and nz: "m1 \ 0" "m2 \ 0" shows "\!x. x < m1 * m2 \ [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - obtain y where y1: "[y = u1] (mod m1)" and y2: "[y = u2] (mod m2)" using binary_chinese_remainder_nat [OF a] by blast let ?x = "y mod (m1 * m2)" from nz have less: "?x < m1 * m2" by auto have 1: "[?x = u1] (mod m1)" using y1 mod_mult_cong_right by blast have 2: "[?x = u2] (mod m2)" using y2 mod_mult_cong_left by blast have "z = ?x" if "z < m1 * m2" "[z = u1] (mod m1)" "[z = u2] (mod m2)" for z proof - have "[?x = z] (mod m1)" by (metis "1" cong_def that(2)) moreover have "[?x = z] (mod m2)" by (metis "2" cong_def that(3)) ultimately have "[?x = z] (mod m1 * m2)" using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) with \z < m1 * m2\ \?x < m1 * m2\ show "z = ?x" by (auto simp add: cong_def) qed with less 1 2 show ?thesis by blast qed lemma chinese_remainder_nat: fixes A :: "'a set" and m :: "'a \ nat" and u :: "'a \ nat" assumes fin: "finite A" and cop: "\i \ A. \j \ A. i \ j \ coprime (m i) (m j)" shows "\x. \i \ A. [x = u i] (mod m i)" proof - have "\b. (\i \ A. [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j)))" proof (rule finite_set_choice, rule fin, rule ballI) fix i assume "i \ A" with cop have "coprime (\j \ A - {i}. m j) (m i)" by (intro prod_coprime_left) auto then have "\x. [(\j \ A - {i}. m j) * x = Suc 0] (mod m i)" by (elim cong_solve_coprime_nat) then obtain x where "[(\j \ A - {i}. m j) * x = 1] (mod m i)" by auto moreover have "[(\j \ A - {i}. m j) * x = 0] (mod (\j \ A - {i}. m j))" by (simp add: cong_0_iff) ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod prod m (A - {i}))" by blast qed then obtain b where b: "\i. i \ A \ [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j))" by blast let ?x = "\i\A. (u i) * (b i)" show ?thesis proof (rule exI, clarify) fix i assume a: "i \ A" show "[?x = u i] (mod m i)" proof - from fin a have "?x = (\j \ {i}. u j * b j) + (\j \ A - {i}. u j * b j)" by (subst sum.union_disjoint [symmetric]) (auto intro: sum.cong) then have "[?x = u i * b i + (\j \ A - {i}. u j * b j)] (mod m i)" by auto also have "[u i * b i + (\j \ A - {i}. u j * b j) = u i * 1 + (\j \ A - {i}. u j * 0)] (mod m i)" proof (intro cong_add cong_scalar_left cong_sum) show "[b i = 1] (mod m i)" using a b by blast show "[b x = 0] (mod m i)" if "x \ A - {i}" for x proof - have "x \ A" "x \ i" using that by auto then show ?thesis using a b [OF \x \ A\] cong_dvd_modulus_nat fin by blast qed qed finally show ?thesis by simp qed qed qed lemma coprime_cong_prod_nat: "[x = y] (mod (\i\A. m i))" if "\i j. \i \ A; j \ A; i \ j\ \ coprime (m i) (m j)" and "\i. i \ A \ [x = y] (mod m i)" for x y :: nat using that proof (induct A rule: infinite_finite_induct) case (insert x A) then show ?case by simp (metis coprime_cong_mult_nat prod_coprime_right) qed auto lemma chinese_remainder_unique_nat: fixes A :: "'a set" and m :: "'a \ nat" and u :: "'a \ nat" assumes fin: "finite A" and nz: "\i\A. m i \ 0" and cop: "\i\A. \j\A. i \ j \ coprime (m i) (m j)" shows "\!x. x < (\i\A. m i) \ (\i\A. [x = u i] (mod m i))" proof - from chinese_remainder_nat [OF fin cop] obtain y where one: "(\i\A. [y = u i] (mod m i))" by blast let ?x = "y mod (\i\A. m i)" from fin nz have prodnz: "(\i\A. m i) \ 0" by auto then have less: "?x < (\i\A. m i)" by auto have cong: "\i\A. [?x = u i] (mod m i)" using fin one by (auto simp add: cong_def dvd_prod_eqI mod_mod_cancel) have unique: "\z. z < (\i\A. m i) \ (\i\A. [z = u i] (mod m i)) \ z = ?x" proof clarify fix z assume zless: "z < (\i\A. m i)" assume zcong: "(\i\A. [z = u i] (mod m i))" have "\i\A. [?x = z] (mod m i)" using cong zcong by (auto simp add: cong_def) with fin cop have "[?x = z] (mod (\i\A. m i))" by (intro coprime_cong_prod_nat) auto with zless less show "z = ?x" by (auto simp add: cong_def) qed from less cong unique show ?thesis by blast qed end diff --git a/src/HOL/Numeral_Simprocs.thy b/src/HOL/Numeral_Simprocs.thy --- a/src/HOL/Numeral_Simprocs.thy +++ b/src/HOL/Numeral_Simprocs.thy @@ -1,299 +1,299 @@ (* Author: Various *) section \Combination and Cancellation Simprocs for Numeral Expressions\ theory Numeral_Simprocs -imports Divides +imports Parity begin ML_file \~~/src/Provers/Arith/assoc_fold.ML\ ML_file \~~/src/Provers/Arith/cancel_numerals.ML\ ML_file \~~/src/Provers/Arith/combine_numerals.ML\ ML_file \~~/src/Provers/Arith/cancel_numeral_factor.ML\ ML_file \~~/src/Provers/Arith/extract_common_term.ML\ lemmas semiring_norm = Let_def arith_simps diff_nat_numeral rel_simps if_False if_True add_Suc add_numeral_left add_neg_numeral_left mult_numeral_left numeral_One [symmetric] uminus_numeral_One [symmetric] Suc_eq_plus1 eq_numeral_iff_iszero not_iszero_Numeral1 text \For \combine_numerals\\ lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" by (simp add: add_mult_distrib) text \For \cancel_numerals\\ lemma nat_diff_add_eq1: "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" by (simp split: nat_diff_split add: add_mult_distrib) lemma nat_diff_add_eq2: "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" by (simp split: nat_diff_split add: add_mult_distrib) lemma nat_eq_add_iff1: "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_eq_add_iff2: "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_less_add_iff1: "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_less_add_iff2: "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_le_add_iff1: "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_le_add_iff2: "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) text \For \cancel_numeral_factors\\ lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" by auto lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" by auto lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" by auto lemma nat_mult_dvd_cancel_disj[simp]: "(k*m) dvd (k*n) = (k=0 \ m dvd (n::nat))" by (auto simp: dvd_eq_mod_eq_0 mod_mult_mult1) lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" by(auto) text \For \cancel_factor\\ lemmas nat_mult_le_cancel_disj = mult_le_cancel1 lemmas nat_mult_less_cancel_disj = mult_less_cancel1 lemma nat_mult_eq_cancel_disj: fixes k m n :: nat shows "k * m = k * n \ k = 0 \ m = n" by (fact mult_cancel_left) lemma nat_mult_div_cancel_disj: fixes k m n :: nat shows "(k * m) div (k * n) = (if k = 0 then 0 else m div n)" by (fact div_mult_mult1_if) lemma numeral_times_minus_swap: fixes x:: "'a::comm_ring_1" shows "numeral w * -x = x * - numeral w" by (simp add: ac_simps) ML_file \Tools/numeral_simprocs.ML\ simproc_setup semiring_assoc_fold ("(a::'a::comm_semiring_1_cancel) * b") = \fn phi => Numeral_Simprocs.assoc_fold\ (* TODO: see whether the type class can be generalized further *) simproc_setup int_combine_numerals ("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") = \fn phi => Numeral_Simprocs.combine_numerals\ simproc_setup field_combine_numerals ("(i::'a::{field,ring_char_0}) + j" |"(i::'a::{field,ring_char_0}) - j") = \fn phi => Numeral_Simprocs.field_combine_numerals\ simproc_setup inteq_cancel_numerals ("(l::'a::comm_ring_1) + m = n" |"(l::'a::comm_ring_1) = m + n" |"(l::'a::comm_ring_1) - m = n" |"(l::'a::comm_ring_1) = m - n" |"(l::'a::comm_ring_1) * m = n" |"(l::'a::comm_ring_1) = m * n" |"- (l::'a::comm_ring_1) = m" |"(l::'a::comm_ring_1) = - m") = \fn phi => Numeral_Simprocs.eq_cancel_numerals\ simproc_setup intless_cancel_numerals ("(l::'a::linordered_idom) + m < n" |"(l::'a::linordered_idom) < m + n" |"(l::'a::linordered_idom) - m < n" |"(l::'a::linordered_idom) < m - n" |"(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n" |"- (l::'a::linordered_idom) < m" |"(l::'a::linordered_idom) < - m") = \fn phi => Numeral_Simprocs.less_cancel_numerals\ simproc_setup intle_cancel_numerals ("(l::'a::linordered_idom) + m \ n" |"(l::'a::linordered_idom) \ m + n" |"(l::'a::linordered_idom) - m \ n" |"(l::'a::linordered_idom) \ m - n" |"(l::'a::linordered_idom) * m \ n" |"(l::'a::linordered_idom) \ m * n" |"- (l::'a::linordered_idom) \ m" |"(l::'a::linordered_idom) \ - m") = \fn phi => Numeral_Simprocs.le_cancel_numerals\ simproc_setup ring_eq_cancel_numeral_factor ("(l::'a::{idom,ring_char_0}) * m = n" |"(l::'a::{idom,ring_char_0}) = m * n") = \fn phi => Numeral_Simprocs.eq_cancel_numeral_factor\ simproc_setup ring_less_cancel_numeral_factor ("(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n") = \fn phi => Numeral_Simprocs.less_cancel_numeral_factor\ simproc_setup ring_le_cancel_numeral_factor ("(l::'a::linordered_idom) * m <= n" |"(l::'a::linordered_idom) <= m * n") = \fn phi => Numeral_Simprocs.le_cancel_numeral_factor\ (* TODO: remove comm_ring_1 constraint if possible *) simproc_setup int_div_cancel_numeral_factors ("((l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) * m) div n" |"(l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) div (m * n)") = \fn phi => Numeral_Simprocs.div_cancel_numeral_factor\ simproc_setup divide_cancel_numeral_factor ("((l::'a::{field,ring_char_0}) * m) / n" |"(l::'a::{field,ring_char_0}) / (m * n)" |"((numeral v)::'a::{field,ring_char_0}) / (numeral w)") = \fn phi => Numeral_Simprocs.divide_cancel_numeral_factor\ simproc_setup ring_eq_cancel_factor ("(l::'a::idom) * m = n" | "(l::'a::idom) = m * n") = \fn phi => Numeral_Simprocs.eq_cancel_factor\ simproc_setup linordered_ring_le_cancel_factor ("(l::'a::linordered_idom) * m <= n" |"(l::'a::linordered_idom) <= m * n") = \fn phi => Numeral_Simprocs.le_cancel_factor\ simproc_setup linordered_ring_less_cancel_factor ("(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n") = \fn phi => Numeral_Simprocs.less_cancel_factor\ simproc_setup int_div_cancel_factor ("((l::'a::euclidean_semiring_cancel) * m) div n" |"(l::'a::euclidean_semiring_cancel) div (m * n)") = \fn phi => Numeral_Simprocs.div_cancel_factor\ simproc_setup int_mod_cancel_factor ("((l::'a::euclidean_semiring_cancel) * m) mod n" |"(l::'a::euclidean_semiring_cancel) mod (m * n)") = \fn phi => Numeral_Simprocs.mod_cancel_factor\ simproc_setup dvd_cancel_factor ("((l::'a::idom) * m) dvd n" |"(l::'a::idom) dvd (m * n)") = \fn phi => Numeral_Simprocs.dvd_cancel_factor\ simproc_setup divide_cancel_factor ("((l::'a::field) * m) / n" |"(l::'a::field) / (m * n)") = \fn phi => Numeral_Simprocs.divide_cancel_factor\ ML_file \Tools/nat_numeral_simprocs.ML\ simproc_setup nat_combine_numerals ("(i::nat) + j" | "Suc (i + j)") = \fn phi => Nat_Numeral_Simprocs.combine_numerals\ simproc_setup nateq_cancel_numerals ("(l::nat) + m = n" | "(l::nat) = m + n" | "(l::nat) * m = n" | "(l::nat) = m * n" | "Suc m = n" | "m = Suc n") = \fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals\ simproc_setup natless_cancel_numerals ("(l::nat) + m < n" | "(l::nat) < m + n" | "(l::nat) * m < n" | "(l::nat) < m * n" | "Suc m < n" | "m < Suc n") = \fn phi => Nat_Numeral_Simprocs.less_cancel_numerals\ simproc_setup natle_cancel_numerals ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "(l::nat) * m \ n" | "(l::nat) \ m * n" | "Suc m \ n" | "m \ Suc n") = \fn phi => Nat_Numeral_Simprocs.le_cancel_numerals\ simproc_setup natdiff_cancel_numerals ("((l::nat) + m) - n" | "(l::nat) - (m + n)" | "(l::nat) * m - n" | "(l::nat) - m * n" | "Suc m - n" | "m - Suc n") = \fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals\ simproc_setup nat_eq_cancel_numeral_factor ("(l::nat) * m = n" | "(l::nat) = m * n") = \fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor\ simproc_setup nat_less_cancel_numeral_factor ("(l::nat) * m < n" | "(l::nat) < m * n") = \fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor\ simproc_setup nat_le_cancel_numeral_factor ("(l::nat) * m <= n" | "(l::nat) <= m * n") = \fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor\ simproc_setup nat_div_cancel_numeral_factor ("((l::nat) * m) div n" | "(l::nat) div (m * n)") = \fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor\ simproc_setup nat_dvd_cancel_numeral_factor ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") = \fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\ simproc_setup nat_eq_cancel_factor ("(l::nat) * m = n" | "(l::nat) = m * n") = \fn phi => Nat_Numeral_Simprocs.eq_cancel_factor\ simproc_setup nat_less_cancel_factor ("(l::nat) * m < n" | "(l::nat) < m * n") = \fn phi => Nat_Numeral_Simprocs.less_cancel_factor\ simproc_setup nat_le_cancel_factor ("(l::nat) * m <= n" | "(l::nat) <= m * n") = \fn phi => Nat_Numeral_Simprocs.le_cancel_factor\ simproc_setup nat_div_cancel_factor ("((l::nat) * m) div n" | "(l::nat) div (m * n)") = \fn phi => Nat_Numeral_Simprocs.div_cancel_factor\ simproc_setup nat_dvd_cancel_factor ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") = \fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor\ declaration \ K (Lin_Arith.add_simprocs [\<^simproc>\semiring_assoc_fold\, \<^simproc>\int_combine_numerals\, \<^simproc>\inteq_cancel_numerals\, \<^simproc>\intless_cancel_numerals\, \<^simproc>\intle_cancel_numerals\, \<^simproc>\field_combine_numerals\, \<^simproc>\nat_combine_numerals\, \<^simproc>\nateq_cancel_numerals\, \<^simproc>\natless_cancel_numerals\, \<^simproc>\natle_cancel_numerals\, \<^simproc>\natdiff_cancel_numerals\, Numeral_Simprocs.field_divide_cancel_numeral_factor]) \ end diff --git a/src/HOL/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2615 +1,2615 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval -imports Divides +imports Parity begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_3_iff: "card S = 3 \ (\x y z. S = {x,y,z} \ x \ y \ y \ z \ x \ z)" by (fastforce simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by (simp add: order_class.order.eq_iff) (auto intro: order_trans) lemma (in linorder) Ico_eq_Ico: "{l.. h=h' \ \ l \ l' a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" proof (cases "a < b") case True assume assm: "{a.. {c.. a \ a \ d" using True by (auto simp add: subset_eq Ball_def) then have 2: "b \ d" using assm by (auto simp add: subset_eq) from 1 2 show ?thesis by simp qed (auto) lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: \{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d\ (is \?P \ ?Q\) proof assume ?Q then show ?P by auto next assume ?P then have \a < x \ x \ b \ c < x \ x \ d\ for x by (simp add: set_eq_iff) from this [of a] this [of b] this [of c] this [of d] show ?Q by auto qed lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (use \y \ -x\[THEN le_imp_neg_le] in \simp\) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. (-) d ` {a..b}" proof fix x assume "x \ {d - b..d - a}" then have "d - x \ {a..b}" and "x = d - (d - x)" by auto then show "x \ (-) d ` {a..b}" by (rule rev_image_eqI) qed qed(auto) lemma image_diff_atLeastLessThan [simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "\f::nat\nat. (!!n. n \ f n) \ finite {n. f n \ u}" by (rule finite_subset[of _ "{..u}"]) (auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" proof (rule UN_finite_subset, rule subsetI) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: assumes "(\n::nat. (\i\{0..i\{0..n. A n) = (\n. B n)" proof (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) fix n show "\ (A ` {0.. (\n. B n)" using assms by auto next fix n show "\ (B ` {0.. \ (A ` {0..Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {.. {l.. (\x. x + l) ` {.. {l.. {..< u -l}" by auto then have "(x - l) + l \ (\x. x + l) ` {..< u -l}" by auto then show "x \ (\x. x + l) ` {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" unfolding inj_on_def proof safe fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed lemma inj_on_funpow_least: \<^marker>\contributor \Lars Noschinski\\ \inj_on (\k. (f ^^ k) s) {0.. if \(f ^^ n) s = s\ \\m. 0 < m \ m < n \ (f ^^ m) s \ s\ proof - { fix k l assume A: "k < n" "l < n" "k \ l" "(f ^^ k) s = (f ^^ l) s" define k' l' where "k' = min k l" and "l' = max k l" with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" by (auto simp: min_def max_def) have "s = (f ^^ ((n - l') + l')) s" using that \l' < n\ by simp also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) finally have "(f ^^ (n - l' + k')) s = s" by simp moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ ultimately have False using that(2) by auto } then show ?thesis by (intro inj_onI) auto qed subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ lemma image_atLeastZeroLessThan_int: assumes "0 \ u" shows "{(0::int).. {y. \x\{x. x < nat u}. y = int x}" proof fix x assume "x \ {0..xa {y. \x\{x. x < nat u}. y = int x}" by simp qed qed (auto) lemma finite_atLeastZeroLessThan_int: "finite {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed lemma card_le_Suc_Max: "finite S \ card S \ Suc (Max S)" proof (rule classical) assume "finite S" and "\ Suc (Max S) \ card S" then have "Suc (Max S) < card S" by simp with \finite S\ have "S \ {0..Max S}" by auto hence "card S \ card {0..Max S}" by (intro card_mono; auto) thus "card S \ Suc (Max S)" by simp qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] by safe (force intro: leI)+ subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" by (rule inj_onI) simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) lemma Nats_infinite: "infinite (\ :: 'a set)" by (metis Nats_def finite_imageD infinite_UNIV_char_0 inj_on_of_nat) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_atMost_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m..Suc n} = F g {m..n}" unfolding atLeast_Suc_atMost_Suc_shift by simp lemma atLeast_lessThan_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma last_plus: fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" by (simp add: commute last_plus) lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma card_sum_le_nat_sum: "\ {0.. \ S" proof (cases "finite S") case True then show ?thesis proof (induction "card S" arbitrary: S) case (Suc x) then have "Max S \ x" using card_le_Suc_Max by fastforce let ?S' = "S - {Max S}" from Suc have "Max S \ S" by (auto intro: Max_in) hence cards: "card S = Suc (card ?S')" using \finite S\ by (intro card.remove; auto) hence "\ {0.. \ ?S'" using Suc by (intro Suc; auto) hence "\ {0.. \ ?S' + Max S" using \Max S \ x\ by simp also have "... = \ S" using sum.remove[OF \finite S\ \Max S \ S\, where g="\x. x"] by simp finally show ?case using cards Suc by auto qed simp qed simp lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxShifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" proof (cases "n < m") case False assume *: "\ n < m" then show ?thesis proof (cases "x = 1") case False assume "x \ 1" then have not_zero: "1 - x \ 0" by auto have "(1 - x) * (\i=m..n. x^i) = x ^ m - x * x ^ n" using sum_gp_multiplied [of m n x] * by auto then have "(\i=m..n. x^i) = (x ^ m - x * x ^ n) / (1 - x) " using nonzero_divide_eq_eq mult.commute not_zero by metis then show ?thesis by auto qed (auto) qed (auto) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end