diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,1185 +1,1187 @@ (* Title: HOL/Divides.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge *) section \More on quotient and remainder\ theory Divides imports Parity begin subsection \More on division\ subsubsection \Splitting Rules for div and mod\ text\The proofs of the two lemmas below are essentially identical\ lemma split_pos_lemma: "0 P(n div k :: int)(n mod k) = (\i j. 0\j \ j n = k*i + j \ P i j)" by auto lemma split_neg_lemma: "k<0 \ P(n div k :: int)(n mod k) = (\i j. k j\0 \ n = k*i + j \ P i j)" by auto lemma split_zdiv: \P (n div k) \ (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))\ for n k :: int proof (cases \k = 0\) case True then show ?thesis by simp next case False then have \k < 0 \ 0 < k\ by auto then show ?thesis by (auto simp add: split_pos_lemma [of concl: "\x y. P x"] split_neg_lemma [of concl: "\x y. P x"]) qed lemma split_zmod: \P (n mod k) \ (k = 0 \ P n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P j))\ for n k :: int proof (cases \k = 0\) case True then show ?thesis by simp next case False then have \k < 0 \ 0 < k\ by auto then show ?thesis by (auto simp add: split_pos_lemma [of concl: "\x y. P y"] split_neg_lemma [of concl: "\x y. P y"]) 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 k\, linarith_split] for k -declare split_zmod [of _ _ \numeral k\, linarith_split] for 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 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 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 \Monotonicity in the First Argument (Dividend)\ inductive eucl_rel_int :: "int \ int \ int \ int \ bool" where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" | eucl_rel_int_dividesI: "l \ 0 \ k = q * l \ eucl_rel_int k l (q, 0)" | eucl_rel_int_remainderI: "sgn r = sgn l \ \r\ < \l\ \ k = q * l + r \ eucl_rel_int k l (q, r)" lemma eucl_rel_int_iff: "eucl_rel_int k l (q, r) \ k = l * q + r \ (if 0 < l then 0 \ r \ r < l else if l < 0 then l < r \ r \ 0 else q = 0)" by (cases "r = 0") (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI simp add: ac_simps sgn_1_pos sgn_1_neg) lemma unique_quotient_lemma: assumes "b * q' + r' \ b * q + r" "0 \ r'" "r' < b" "r < b" shows "q' \ (q::int)" proof - have "r' + b * (q'-q) \ r" using assms by (simp add: right_diff_distrib) moreover have "0 < b * (1 + q - q') " using assms by (simp add: right_diff_distrib distrib_left) moreover have "b * q' < b * (1 + q)" using assms by (simp add: right_diff_distrib distrib_left) ultimately show ?thesis using assms by (simp add: mult_less_cancel_left) qed lemma unique_quotient_lemma_neg: "b * q' + r' \ b*q + r \ r \ 0 \ b < r \ b < r' \ q \ (q'::int)" using unique_quotient_lemma[where b = "-b" and r = "-r'" and r'="-r"] by auto lemma unique_quotient: "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ q = q'" apply (rule order_antisym) apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ done lemma unique_remainder: assumes "eucl_rel_int a b (q, r)" and "eucl_rel_int a b (q', r')" shows "r = r'" proof - have "q = q'" using assms by (blast intro: unique_quotient) then show "r = r'" using assms by (simp add: eucl_rel_int_iff) qed lemma eucl_rel_int: "eucl_rel_int k l (k div l, k mod l)" proof (cases k rule: int_cases3) case zero then show ?thesis by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def) next case (pos n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) next case (neg n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) qed lemma divmod_int_unique: assumes "eucl_rel_int k l (q, r)" shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" using assms eucl_rel_int [of k l] using unique_quotient [of k l] unique_remainder [of k l] by auto lemma zminus1_lemma: "eucl_rel_int a b (q, r) ==> b \ 0 ==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1, if r=0 then 0 else b-r)" by (force simp add: eucl_rel_int_iff right_diff_distrib) lemma zdiv_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) lemma div_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k div l = (k - l) div l + 1" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by (simp add: div_add_self2) qed lemma mod_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k mod l = (k - l) mod l" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by simp qed subsubsection \Computing \div\ and \mod\ with shifting\ lemma pos_eucl_rel_int_mult_2: assumes "0 \ b" assumes "eucl_rel_int a b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)" using assms unfolding eucl_rel_int_iff by auto lemma neg_eucl_rel_int_mult_2: assumes "b \ 0" assumes "eucl_rel_int (a + 1) b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)" using assms unfolding eucl_rel_int_iff by auto text\computing div by shifting\ lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int] by (rule div_int_unique) lemma neg_zdiv_mult_2: assumes A: "a \ (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int] by (rule div_int_unique) lemma zdiv_numeral_Bit0 [simp]: "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = numeral v div (numeral w :: int)" unfolding numeral.simps unfolding mult_2 [symmetric] by (rule div_mult_mult1, simp) lemma zdiv_numeral_Bit1 [simp]: "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = (numeral v div (numeral w :: int))" unfolding numeral.simps unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zdiv_mult_2, simp) lemma pos_zmod_mult_2: fixes a b :: int assumes "0 \ a" shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) lemma neg_zmod_mult_2: fixes a b :: int assumes "a \ 0" shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) lemma zmod_numeral_Bit0 [simp]: "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = (2::int) * (numeral v mod numeral w)" unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] by (rule mod_mult_mult1) lemma zmod_numeral_Bit1 [simp]: "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = 2 * (numeral v mod numeral w) + (1::int)" unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zmod_mult_2, simp) 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 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 int_div_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a div b = q" by (rule div_int_unique [of a b q r]) (simp add: eucl_rel_int_iff) lemma int_div_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a div b = q" by (rule div_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma abs_div: "(y::int) dvd x \ \x div y\ = \x\ div \y\" unfolding dvd_def by (cases "y=0") (auto simp add: abs_mult) text\Suggested by Matthias Daum\ lemma int_power_div_base: fixes k :: int assumes "0 < m" "0 < k" shows "k ^ m div k = (k::int) ^ (m - Suc 0)" proof - have eq: "k ^ m = k ^ ((m - Suc 0) + Suc 0)" by (simp add: assms) show ?thesis using assms by (simp only: power_add eq) auto qed text\Suggested by Matthias Daum\ lemma int_div_less_self: fixes x::int assumes "0 < x" "1 < k" shows "x div k < x" proof - have "nat x div nat k < nat x" by (simp add: assms) with assms show ?thesis by (simp add: nat_div_distrib [symmetric]) qed lemma mod_eq_dvd_iff_nat: "m mod q = n mod q \ q dvd m - n" if "m \ n" for m n q :: nat proof - have "int m mod int q = int n mod int q \ int q dvd int m - int n" by (simp add: mod_eq_dvd_iff) with that have "int (m mod q) = int (n mod q) \ int q dvd int (m - n)" by (simp only: of_nat_mod of_nat_diff) then show ?thesis by simp qed lemma mod_eq_nat1E: fixes m n q :: nat assumes "m mod q = n mod q" and "m \ n" obtains s where "m = n + q * s" proof - from assms have "q dvd m - n" by (simp add: mod_eq_dvd_iff_nat) then obtain s where "m - n = q * s" .. with \m \ n\ have "m = n + q * s" by simp with that show thesis . qed lemma mod_eq_nat2E: fixes m n q :: nat assumes "m mod q = n mod q" and "n \ m" obtains s where "n = m + q * s" using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) lemma nat_mod_eq_lemma: assumes "(x::nat) mod n = y mod n" and "y \ x" shows "\q. x = y + n * q" using assms by (rule mod_eq_nat1E) (rule 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 subsection \Numeral division with a pragmatic type class\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. This is its primary motivation, and it could surely be formulated using a more fine-grained, more algebraic and less technical class hierarchy. \ class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom + assumes div_less: "0 \ a \ a < b \ a div b = 0" and mod_less: " 0 \ a \ a < b \ a mod b = a" and div_positive: "0 < b \ b \ a \ a div b > 0" and mod_less_eq_dividend: "0 \ a \ a mod b \ a" and pos_mod_bound: "0 < b \ a mod b < b" and pos_mod_sign: "0 < b \ 0 \ a mod b" and mod_mult2_eq: "0 \ c \ a mod (b * c) = b * (a div b mod c) + a mod b" and div_mult2_eq: "0 \ c \ a div (b * c) = a div b div c" assumes discrete: "a < b \ a + 1 \ b" fixes divmod :: "num \ num \ 'a \ 'a" and divmod_step :: "num \ 'a \ 'a \ 'a \ 'a" assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)" and divmod_step_def: "divmod_step l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" \ \These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ begin lemma divmod_digit_1: assumes "0 \ a" "0 < b" and "b \ a mod (2 * b)" shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") and "a mod (2 * b) - b = a mod b" (is "?Q") proof - from assms mod_less_eq_dividend [of a "2 * b"] have "b \ a" by (auto intro: trans) with \0 < b\ have "0 < a div b" by (auto intro: div_positive) then have [simp]: "1 \ a div b" by (simp add: discrete) with \0 < b\ have mod_less: "a mod b < b" by (simp add: pos_mod_bound) define w where "w = a div b mod 2" then have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) from assms w_exhaust have "w = 1" 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 lemma fst_divmod: "fst (divmod m n) = numeral m div numeral n" by (simp add: divmod_def) lemma snd_divmod: "snd (divmod m n) = numeral m mod numeral n" by (simp add: divmod_def) text \ This is a formulation of one step (referring to one digit position) in school-method division: compare the dividend at the current digit position with the remainder from previous division steps and evaluate accordingly. \ lemma divmod_step_eq [simp]: "divmod_step l (q, r) = (if numeral l \ r then (2 * q + 1, r - numeral l) else (2 * q, r))" by (simp add: divmod_step_def) text \ This is a formulation of school-method division. If the divisor is smaller than the dividend, terminate. If not, shift the dividend to the right until termination occurs and then reiterate single division steps in the opposite direction. \ lemma divmod_divmod_step: "divmod m n = (if m < n then (0, numeral m) else divmod_step n (divmod m (Num.Bit0 n)))" proof (cases "m < n") case True then have "numeral m < numeral n" by simp then show ?thesis by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod) next case False have "divmod m n = divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n))" proof (cases "numeral n \ numeral m mod (2 * numeral n)") case True with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)" by simp moreover from True divmod_digit_1 [of "numeral m" "numeral n"] have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n" and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n" by simp_all ultimately show ?thesis by (simp only: divmod_def) next case False then have *: "numeral m mod (2 * numeral n) < numeral n" by (simp add: not_le) with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))" by auto moreover from * divmod_digit_0 [of "numeral n" "numeral m"] have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n" and "numeral m mod (2 * numeral n) = numeral m mod numeral n" by (simp_all only: zero_less_numeral) ultimately show ?thesis by (simp only: divmod_def) qed then have "divmod m n = divmod_step n (numeral m div numeral (Num.Bit0 n), numeral m mod numeral (Num.Bit0 n))" by (simp only: numeral.simps distrib mult_1) then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))" by (simp add: divmod_def) with False show ?thesis by simp qed text \The division rewrite proper -- first, trivial results involving \1\\ lemma divmod_trivial [simp]: "divmod m Num.One = (numeral m, 0)" "divmod num.One (num.Bit0 n) = (0, Numeral1)" "divmod num.One (num.Bit1 n) = (0, Numeral1)" using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) text \Division by an even number is a right-shift\ lemma divmod_cancel [simp]: "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r))" (is ?P) "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r + 1))" (is ?Q) proof - have *: "\q. numeral (Num.Bit0 q) = 2 * numeral q" "\q. numeral (Num.Bit1 q) = 2 * numeral q + 1" by (simp_all only: numeral_mult numeral.simps distrib) simp_all have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less) then show ?P and ?Q by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1 div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] add.commute del: numeral_times_numeral) qed text \The really hard work\ lemma divmod_steps [simp]: "divmod (num.Bit0 m) (num.Bit1 n) = (if m \ n then (0, numeral (num.Bit0 m)) else divmod_step (num.Bit1 n) (divmod (num.Bit0 m) (num.Bit0 (num.Bit1 n))))" "divmod (num.Bit1 m) (num.Bit1 n) = (if m < n then (0, numeral (num.Bit1 m)) else divmod_step (num.Bit1 n) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps text \Special case: divisibility\ definition divides_aux :: "'a \ 'a \ bool" where "divides_aux qr \ snd qr = 0" lemma divides_aux_eq [simp]: "divides_aux (q, r) \ r = 0" by (simp add: divides_aux_def) lemma dvd_numeral_simp [simp]: "numeral m dvd numeral n \ divides_aux (divmod n m)" by (simp add: divmod_def mod_eq_0_iff_dvd) text \Generic computation of quotient and remainder\ lemma numeral_div_numeral [simp]: "numeral k div numeral l = fst (divmod k l)" by (simp add: fst_divmod) lemma numeral_mod_numeral [simp]: "numeral k mod numeral l = snd (divmod k l)" by (simp add: snd_divmod) lemma one_div_numeral [simp]: "1 div numeral n = fst (divmod num.One n)" by (simp add: fst_divmod) lemma one_mod_numeral [simp]: "1 mod numeral n = snd (divmod num.One n)" by (simp add: snd_divmod) text \Computing congruences modulo \2 ^ q\\ lemma cong_exp_iff_simps: "numeral n mod numeral Num.One = 0 \ True" "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 \ numeral n mod numeral q = 0" "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 \ False" "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ (numeral n mod numeral q) = 0" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ (numeral m mod numeral q) = 0" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) end hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq instantiation nat :: unique_euclidean_semiring_numeral begin definition divmod_nat :: "num \ num \ nat \ nat" where divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_nat :: "num \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: fixes k l :: num shows "Suc 0 div numeral k = fst (divmod Num.One k)" by (simp_all add: fst_divmod) lemma Suc_0_mod_numeral [simp]: fixes k l :: num shows "Suc 0 mod numeral k = snd (divmod Num.One k)" by (simp_all add: snd_divmod) instantiation int :: unique_euclidean_semiring_numeral begin definition divmod_int :: "num \ num \ int \ int" where "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_int :: "num \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) end declare divmod_algorithm_code [where ?'a = int, code] context begin qualified definition adjust_div :: "int \ int \ int" where "adjust_div qr = (let (q, r) = qr in q + of_bool (r \ 0))" qualified lemma adjust_div_eq [simp, code]: "adjust_div (q, r) = q + of_bool (r \ 0)" by (simp add: adjust_div_def) qualified definition adjust_mod :: "int \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else l - r)" lemma minus_numeral_div_numeral [simp]: "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma minus_numeral_mod_numeral [simp]: "- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma numeral_div_minus_numeral [simp]: "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma numeral_mod_minus_numeral [simp]: "numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma minus_one_div_numeral [simp]: "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" using minus_numeral_div_numeral [of Num.One n] by simp lemma minus_one_mod_numeral [simp]: "- 1 mod numeral n = adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" using minus_numeral_mod_numeral [of Num.One n] by simp lemma one_div_minus_numeral [simp]: "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" using numeral_div_minus_numeral [of Num.One n] by simp lemma one_mod_minus_numeral [simp]: "1 mod - numeral n = - adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp end lemma divmod_BitM_2_eq [simp]: \divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\ by (cases m) simp_all lemma div_positive_int: "k div l > 0" if "k \ l" and "l > 0" for k l :: int using that div_positive [of l k] by blast subsubsection \Dedicated simproc for calculation\ text \ There is space for improvement here: the calculation itself could be carried out outside the logic, and a generic simproc (simplifier setup) for generic calculation would be helpful. \ simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | "numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \ let val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\If\); fun successful_rewrite ctxt ct = let val thm = Simplifier.rewrite ctxt ct in if Thm.is_reflexive thm then NONE else SOME thm end; in fn phi => let val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral one_div_minus_numeral one_mod_minus_numeral numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral numeral_div_minus_numeral numeral_mod_minus_numeral div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right} @ [@{lemma "0 = 0 \ True" by simp}]); fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end end \ subsubsection \Code generation\ definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" lemma fst_divmod_nat [simp]: "fst (divmod_nat m n) = m div n" by (simp add: divmod_nat_def) lemma snd_divmod_nat [simp]: "snd (divmod_nat m n) = m mod n" by (simp add: divmod_nat_def) lemma divmod_nat_if [code]: "Divides.divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))" by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by simp_all lemma [code]: fixes k :: int shows "k div 0 = 0" "k mod 0 = k" "0 div k = 0" "0 mod k = 0" "k div Int.Pos Num.One = k" "k mod Int.Pos Num.One = 0" "k div Int.Neg Num.One = - k" "k mod Int.Neg Num.One = 0" "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" "Int.Neg m div Int.Pos n = - (Divides.adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" by simp_all code_identifier code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Lemmas of doubtful value\ lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \ m < n" for m n :: nat by (rule le_div_geq) (use that in \simp_all add: not_less\) lemma mod_geq: "m mod n = (m - n) mod n" if "\ m < n" for m n :: nat by (rule le_mod_geq) (use that in \simp add: not_less\) lemma mod_eq_0D: "\q. m = d * q" if "m mod d = 0" for m d :: nat using that by (auto simp add: mod_eq_0_iff_dvd) lemma pos_mod_conj: "0 < b \ 0 \ a mod b \ a mod b < b" for a b :: int by simp lemma neg_mod_conj: "b < 0 \ a mod b \ 0 \ b < a mod b" for a b :: int by simp lemma zmod_eq_0_iff: "m mod d = 0 \ (\q. m = d * q)" for m d :: int by (auto simp add: mod_eq_0_iff_dvd) (* REVISIT: should this be generalized to all semiring_div types? *) lemma zmod_eq_0D [dest!]: "\q. m = d * q" if "m mod d = 0" for m d :: int using that by auto end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,2259 +1,2262 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (use assms in simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule iffI; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") auto lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ lemma power_mod [mod_simps]: "((a mod b) ^ n) mod b = (a ^ n) mod b" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" by (simp add: mod_mult_right_eq) with Suc show ?case by (simp add: mod_mult_left_eq mod_mult_right_eq) qed lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp qed end class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel begin subclass idom_divide .. lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" using div_mult_mult1 [of "- 1" a b] by simp lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" using mod_mult_mult1 [of "- 1" a b] by simp lemma div_minus_right: "a div (- b) = (- a) div b" using div_minus_minus [of "- a" b] by simp lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" using mod_minus_minus [of "- a" b] by simp lemma div_minus1_right [simp]: "a div (- 1) = - a" using div_minus_right [of a 1] by simp lemma mod_minus1_right [simp]: "a mod (- 1) = 0" using mod_minus_right [of a 1] by simp text \Negation respects modular equivalence.\ lemma mod_minus_eq [mod_simps]: "(- (a mod b)) mod b = (- a) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: div_mult_mod_eq) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. then show ?thesis by (simp add: mod_minus_eq) qed text \Subtraction respects modular equivalence.\ lemma mod_diff_left_eq [mod_simps]: "(a mod c - b) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp lemma mod_diff_right_eq [mod_simps]: "(a - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_eq: "(a mod c - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp lemma minus_mod_self2 [simp]: "(a - b) mod b = a mod b" using mod_diff_right_eq [of a b b] by (simp add: mod_diff_right_eq) lemma minus_mod_self1 [simp]: "(b - a) mod b = - a mod b" using mod_add_self2 [of "- a" b] by simp lemma mod_eq_dvd_iff: "a mod c = b mod c \ c dvd a - b" (is "?P \ ?Q") proof assume ?P then have "(a mod c - b mod c) mod c = 0" by simp then show ?Q by (simp add: dvd_eq_mod_eq_0 mod_simps) next assume ?Q then obtain d where d: "a - b = c * d" .. then have "a = c * d + b" by (simp add: algebra_simps) then show ?P by simp qed lemma mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: "euclidean_size (a * b) = euclidean_size a * euclidean_size b" fixes division_segment :: "'a \ 'a" assumes is_unit_division_segment [simp]: "is_unit (division_segment a)" and division_segment_mult: "a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b" and division_segment_mod: "b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b" assumes div_bounded: "b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q" begin lemma division_segment_not_0 [simp]: "division_segment a \ 0" using is_unit_division_segment [of a] is_unitE [of "division_segment a"] by blast lemma divmod_cases [case_names divides remainder by0]: obtains (divides) q where "b \ 0" and "a div b = q" and "a mod b = 0" and "a = q * b" | (remainder) q r where "b \ 0" and "division_segment r = division_segment b" and "euclidean_size r < euclidean_size b" and "r \ 0" and "a div b = q" and "a mod b = r" and "a = q * b + r" | (by0) "b = 0" proof (cases "b = 0") case True then show thesis by (rule by0) next case False show thesis proof (cases "b dvd a") case True then obtain q where "a = b * q" .. with \b \ 0\ divides show thesis by (simp add: ac_simps) next case False then have "a mod b \ 0" by (simp add: mod_eq_0_iff_dvd) moreover from \b \ 0\ \\ b dvd a\ have "division_segment (a mod b) = division_segment b" by (rule division_segment_mod) moreover have "euclidean_size (a mod b) < euclidean_size b" using \b \ 0\ by (rule mod_size_less) moreover have "a = a div b * b + a mod b" by (simp add: div_mult_mod_eq) ultimately show thesis using \b \ 0\ by (blast intro!: remainder) qed qed lemma div_eqI: "a div b = q" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "(q * b + r) div b = q" by (auto intro: div_bounded) with that show ?thesis by simp qed lemma mod_eqI: "a mod b = r" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "a div b = q" by (rule div_eqI) moreover have "a div b * b + a mod b = a" by (fact div_mult_mod_eq) ultimately have "a div b * b + a mod b = a div b * b + r" using \q * b + r = a\ by simp then show ?thesis by simp qed subclass euclidean_semiring_cancel proof show "(a + c * b) div b = c + a div b" if "b \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 with \b \ 0\ show ?thesis by simp next case (divides q) then show ?thesis by (simp add: ac_simps) next case (remainder q r) then show ?thesis by (auto intro: div_eqI simp add: algebra_simps) qed next show"(c * a) div (c * b) = a div b" if "c \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 then show ?thesis by simp next case (divides q) with \c \ 0\ show ?thesis by (simp add: mult.left_commute [of c]) next case (remainder q r) from \b \ 0\ \c \ 0\ have "b * c \ 0" by simp from remainder \c \ 0\ have "division_segment (r * c) = division_segment (b * c)" and "euclidean_size (r * c) < euclidean_size (b * c)" by (simp_all add: division_segment_mult division_segment_mod euclidean_size_mult) with remainder show ?thesis by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps) (use \b * c \ 0\ in simp) qed qed lemma div_mult1_eq: "(a * b) div c = a * (b div c) + a * (b mod c) div c" proof (cases "a * (b mod c)" c rule: divmod_cases) case (divides q) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = (a * (b div c) + q) * c" using divides by (simp add: algebra_simps) finally have "(a * b) div c = \ div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = a * c * (b div c) + q * c + r" using remainder by (simp add: algebra_simps) finally show "(a * (b div c) + a * (b mod c) div c) * c + r = a * b" using remainder(5-7) by (simp add: algebra_simps) qed next case by0 then show ?thesis by simp qed lemma div_add1_eq: "(a + b) div c = a div c + b div c + (a mod c + b mod c) div c" proof (cases "a mod c + b mod c" c rule: divmod_cases) case (divides q) have "a + b = (a div c * c + a mod c) + (b div c * c + b mod c)" using mod_mult_div_eq [of a c] mod_mult_div_eq [of b c] by (simp add: ac_simps) also have "\ = (a div c + b div c) * c + (a mod c + b mod c)" by (simp add: algebra_simps) also have "\ = (a div c + b div c + q) * c" using divides by (simp add: algebra_simps) finally have "(a + b) div c = (a div c + b div c + q) * c div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "(a div c + b div c + q) * c + r + (a mod c + b mod c) = (a div c * c + a mod c) + (b div c * c + b mod c) + q * c + r" by (simp add: algebra_simps) also have "\ = a + b + (a mod c + b mod c)" by (simp add: div_mult_mod_eq remainder) (simp add: ac_simps) finally show "(a div c + b div c + (a mod c + b mod c) div c) * c + r = a + b" using remainder by simp qed next case by0 then show ?thesis by simp qed lemma div_eq_0_iff: "a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0" (is "_ \ ?P") if "division_segment a = division_segment b" proof assume ?P with that show "a div b = 0" by (cases "b = 0") (auto intro: div_eqI) next assume "a div b = 0" then have "a mod b = a" using div_mult_mod_eq [of a b] by simp with mod_size_less [of b a] show ?P by auto qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: "nat \ nat" where [simp]: "normalize = (id :: nat \ nat)" definition unit_factor_nat :: "nat \ nat" where "unit_factor n = (if n = 0 then 0 else 1 :: nat)" lemma unit_factor_simps [simp]: "unit_factor 0 = (0::nat)" "unit_factor (Suc n) = 1" by (simp_all add: unit_factor_nat_def) definition divide_nat :: "nat \ nat \ nat" where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \ m})" instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: "nat \ nat" where [simp]: "euclidean_size_nat = id" definition division_segment_nat :: "nat \ nat" where [simp]: "division_segment_nat n = 1" definition modulo_nat :: "nat \ nat \ nat" where "m mod n = m - (m div n * (n::nat))" instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_nat_eqI: "m div n = q" if "n * q \ m" and "m < n * Suc q" for m n q :: nat by (rule div_eqI [of _ "m - n * q"]) (use that in \simp_all add: algebra_simps\) lemma mod_nat_eqI: "m mod n = r" if "r < n" and "r \ m" and "n dvd m - r" for m n r :: nat by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \simp_all add: algebra_simps\) lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_eqI mod_eqI) lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all context fixes m n q :: nat begin private lemma eucl_rel_mult2: "m mod n + n * (m div n mod q) < n * q" if "n > 0" and "q > 0" proof - from \n > 0\ have "m mod n < n" by (rule mod_less_divisor) from \q > 0\ have "m div n mod q < q" by (rule mod_less_divisor) then obtain s where "q = Suc (m div n mod q + s)" by (blast dest: less_imp_Suc_add) moreover have "m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)" using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately show ?thesis by simp qed lemma div_mult2_eq: "m div (n * q) = (m div n) div q" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: div_eqI [of _ "n * (m div n mod q) + m mod n"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed lemma mod_mult2_eq: "m mod (n * q) = n * (m div n mod q) + m mod n" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: mod_eqI [of _ _ "(m div n) div q"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed end lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed 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) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma split_div: "P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))" (is "?P = ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P i" by simp with False show ?P by (auto intro: * [of "m mod n"]) qed qed lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma split_mod: "P (m mod n) \ (n = 0 \ P m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P j))" (is "?P \ ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P j" by simp with False show ?P by (auto intro: * [of _ "m div n"]) qed qed +declare split_div [of _ _ \numeral n\, linarith_split] for n +declare split_mod [of _ _ \numeral n\, linarith_split] for n + 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 subsection \Elementary euclidean division on \<^typ>\int\\ subsubsection \Basic instantiation\ instantiation int :: "{normalization_semidom, idom_modulo}" begin definition normalize_int :: \int \ int\ where [simp]: \normalize = (abs :: int \ int)\ definition unit_factor_int :: \int \ int\ where [simp]: \unit_factor = (sgn :: int \ int)\ definition divide_int :: \int \ int \ int\ where \k div l = (sgn k * sgn l * int (nat \k\ div nat \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k))\ lemma divide_int_unfold: \(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n) - of_bool ((k = 0 \ m = 0) \ l \ 0 \ n \ 0 \ sgn k \ sgn l \ \ n dvd m))\ by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps) definition modulo_int :: \int \ int \ int\ where \k mod l = sgn k * int (nat \k\ mod nat \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ lemma modulo_int_unfold: \(sgn k * int m) mod (sgn l * int n) = sgn k * int (m mod (of_bool (l \ 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \ m = 0) \ sgn k \ sgn l \ \ n dvd m)\ by (auto simp add: modulo_int_def sgn_mult abs_mult) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) next fix k l :: int obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end subsubsection \Algebraic foundations\ lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all subsubsection \Basic conversions\ lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (auto simp add: divide_int_def) lemma div_eq_div_abs: \k div l = sgn k * sgn l * (\k\ div \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: divide_int_def [of k l] div_abs_eq_div_nat) lemma div_abs_eq: \\k\ div \l\ = sgn k * sgn l * (k div l + of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (simp add: div_eq_div_abs [of k l] ac_simps) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma mod_eq_mod_abs: \k mod l = sgn k * (\k\ mod \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat) lemma mod_abs_eq: \\k\ mod \l\ = sgn k * (k mod l - l * of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (auto simp: mod_eq_mod_abs [of k l]) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" using assms by (simp add: sgn_mult abs_mult sgn_0_0 divide_int_def [of "sgn v * \k\" "sgn v * \l\"] flip: div_abs_eq_div_nat) lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" using assms by (auto simp add: div_abs_eq) lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" using assms by (auto simp add: div_abs_eq ac_simps) lemma div_noneq_sgn_abs: fixes k l :: int assumes "l \ 0" assumes "sgn k \ sgn l" shows "k div l = - (\k\ div \l\) - of_bool (\ l dvd k)" using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp) subsubsection \Euclidean division\ instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) (simp add: sgn_0_0) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: * q l divide_int_unfold) (auto simp add: sgn_mult ac_simps) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) end subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" proof (cases a "of_nat m * of_nat n" rule: divmod_cases) case (divides q) then show ?thesis using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] by (simp add: ac_simps) next case (remainder q r) then have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by simp with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with remainder(6) have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with remainder(1) have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp with remainder(5) remainder(7) show ?thesis using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] by (simp add: ac_simps) next case by0 then show ?thesis by auto qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD) 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 = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule div_eqI [of _ "k + l"]) using False that apply (simp_all add: division_segment_int_def) done qed lemma mod_pos_neg_trivial: "k mod l = k + l" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule mod_eqI [of _ _ \- 1\]) using False that apply (simp_all add: division_segment_int_def) done qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ 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 \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 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) subsection \Code generation\ code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Fields.thy b/src/HOL/Fields.thy --- a/src/HOL/Fields.thy +++ b/src/HOL/Fields.thy @@ -1,1348 +1,1322 @@ (* Title: HOL/Fields.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Fields\ theory Fields imports Nat begin -context idom -begin - -lemma inj_mult_left [simp]: \inj ((*) a) \ a \ 0\ (is \?P \ ?Q\) -proof - assume ?P - show ?Q - proof - assume \a = 0\ - with \?P\ have "inj ((*) 0)" - by simp - moreover have "0 * 0 = 0 * 1" - by simp - ultimately have "0 = 1" - by (rule injD) - then show False - by simp - qed -next - assume ?Q then show ?P - by (auto intro: injI) -qed - -end - - subsection \Division rings\ text \ A division ring is like a field, but without the commutativity requirement. \ class inverse = divide + fixes inverse :: "'a \ 'a" begin abbreviation inverse_divide :: "'a \ 'a \ 'a" (infixl "'/" 70) where "inverse_divide \ divide" end text \Setup for linear arithmetic prover\ ML_file \~~/src/Provers/Arith/fast_lin_arith.ML\ ML_file \Tools/lin_arith.ML\ setup \Lin_Arith.global_setup\ -declaration \K ( +declaration \K ( Lin_Arith.init_arith_data #> Lin_Arith.add_discrete_type \<^type_name>\nat\ #> Lin_Arith.add_lessD @{thm Suc_leI} #> Lin_Arith.add_simps @{thms simp_thms ring_distribs if_True if_False minus_diff_eq add_0_left add_0_right order_less_irrefl zero_neq_one zero_less_one zero_le_one zero_neq_one [THEN not_sym] not_one_le_zero not_one_less_zero add_Suc add_Suc_right nat.inject Suc_le_mono Suc_less_eq Zero_not_Suc Suc_not_Zero le_0_eq One_nat_def} #> Lin_Arith.add_simprocs [\<^simproc>\group_cancel_add\, \<^simproc>\group_cancel_diff\, \<^simproc>\group_cancel_eq\, \<^simproc>\group_cancel_le\, \<^simproc>\group_cancel_less\, \<^simproc>\nateq_cancel_sums\,\<^simproc>\natless_cancel_sums\, \<^simproc>\natle_cancel_sums\])\ simproc_setup fast_arith_nat ("(m::nat) < n" | "(m::nat) \ n" | "(m::nat) = n") = \K Lin_Arith.simproc\ \ \Because of this simproc, the arithmetic solver is really only useful to detect inconsistencies among the premises for subgoals which are \<^emph>\not\ themselves (in)equalities, because the latter activate \<^text>\fast_nat_arith_simproc\ anyway. However, it seems cheaper to activate the solver all the time rather than add the additional check.\ -lemmas [linarith_split] = nat_diff_split split_min split_max +lemmas [linarith_split] = nat_diff_split split_min split_max abs_split text\Lemmas \divide_simps\ move division to the outside and eliminates them on (in)equalities.\ named_theorems divide_simps "rewrite rules to eliminate divisions" class division_ring = ring_1 + inverse + assumes left_inverse [simp]: "a \ 0 \ inverse a * a = 1" assumes right_inverse [simp]: "a \ 0 \ a * inverse a = 1" assumes divide_inverse: "a / b = a * inverse b" assumes inverse_zero [simp]: "inverse 0 = 0" begin subclass ring_1_no_zero_divisors proof fix a b :: 'a assume a: "a \ 0" and b: "b \ 0" show "a * b \ 0" proof assume ab: "a * b = 0" hence "0 = inverse a * (a * b) * inverse b" by simp also have "\ = (inverse a * a) * (b * inverse b)" by (simp only: mult.assoc) also have "\ = 1" using a b by simp finally show False by simp qed qed lemma nonzero_imp_inverse_nonzero: "a \ 0 \ inverse a \ 0" proof assume ianz: "inverse a = 0" assume "a \ 0" hence "1 = a * inverse a" by simp also have "... = 0" by (simp add: ianz) finally have "1 = 0" . thus False by (simp add: eq_commute) qed lemma inverse_zero_imp_zero: assumes "inverse a = 0" shows "a = 0" proof (rule ccontr) assume "a \ 0" then have "inverse a \ 0" by (simp add: nonzero_imp_inverse_nonzero) with assms show False by auto qed lemma inverse_unique: assumes ab: "a * b = 1" shows "inverse a = b" proof - have "a \ 0" using ab by (cases "a = 0") simp_all moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) ultimately show ?thesis by (simp add: mult.assoc [symmetric]) qed lemma nonzero_inverse_minus_eq: "a \ 0 \ inverse (- a) = - inverse a" by (rule inverse_unique) simp lemma nonzero_inverse_inverse_eq: "a \ 0 \ inverse (inverse a) = a" by (rule inverse_unique) simp lemma nonzero_inverse_eq_imp_eq: assumes "inverse a = inverse b" and "a \ 0" and "b \ 0" shows "a = b" proof - from \inverse a = inverse b\ have "inverse (inverse a) = inverse (inverse b)" by (rule arg_cong) with \a \ 0\ and \b \ 0\ show "a = b" by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_1 [simp]: "inverse 1 = 1" by (rule inverse_unique) simp lemma nonzero_inverse_mult_distrib: assumes "a \ 0" and "b \ 0" shows "inverse (a * b) = inverse b * inverse a" proof - have "a * (b * inverse b) * inverse a = 1" using assms by simp hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult.assoc) thus ?thesis by (rule inverse_unique) qed lemma division_ring_inverse_add: "a \ 0 \ b \ 0 \ inverse a + inverse b = inverse a * (a + b) * inverse b" by (simp add: algebra_simps) lemma division_ring_inverse_diff: "a \ 0 \ b \ 0 \ inverse a - inverse b = inverse a * (b - a) * inverse b" by (simp add: algebra_simps) lemma right_inverse_eq: "b \ 0 \ a / b = 1 \ a = b" proof assume neq: "b \ 0" { hence "a = (a / b) * b" by (simp add: divide_inverse mult.assoc) also assume "a / b = 1" finally show "a = b" by simp next assume "a = b" with neq show "a / b = 1" by (simp add: divide_inverse) } qed lemma nonzero_inverse_eq_divide: "a \ 0 \ inverse a = 1 / a" by (simp add: divide_inverse) lemma divide_self [simp]: "a \ 0 \ a / a = 1" by (simp add: divide_inverse) lemma inverse_eq_divide [field_simps, field_split_simps, divide_simps]: "inverse a = 1 / a" by (simp add: divide_inverse) lemma add_divide_distrib: "(a+b) / c = a/c + b/c" by (simp add: divide_inverse algebra_simps) lemma times_divide_eq_right [simp]: "a * (b / c) = (a * b) / c" by (simp add: divide_inverse mult.assoc) lemma minus_divide_left: "- (a / b) = (-a) / b" by (simp add: divide_inverse) lemma nonzero_minus_divide_right: "b \ 0 \ - (a / b) = a / (- b)" by (simp add: divide_inverse nonzero_inverse_minus_eq) lemma nonzero_minus_divide_divide: "b \ 0 \ (-a) / (-b) = a / b" by (simp add: divide_inverse nonzero_inverse_minus_eq) lemma divide_minus_left [simp]: "(-a) / b = - (a / b)" by (simp add: divide_inverse) lemma diff_divide_distrib: "(a - b) / c = a / c - b / c" using add_divide_distrib [of a "- b" c] by simp lemma nonzero_eq_divide_eq [field_simps]: "c \ 0 \ a = b / c \ a * c = b" proof - assume [simp]: "c \ 0" have "a = b / c \ a * c = (b / c) * c" by simp also have "... \ a * c = b" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed lemma nonzero_divide_eq_eq [field_simps]: "c \ 0 \ b / c = a \ b = a * c" proof - assume [simp]: "c \ 0" have "b / c = a \ (b / c) * c = a * c" by simp also have "... \ b = a * c" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed lemma nonzero_neg_divide_eq_eq [field_simps]: "b \ 0 \ - (a / b) = c \ - a = c * b" using nonzero_divide_eq_eq[of b "-a" c] by simp lemma nonzero_neg_divide_eq_eq2 [field_simps]: "b \ 0 \ c = - (a / b) \ c * b = - a" using nonzero_neg_divide_eq_eq[of b a c] by auto lemma divide_eq_imp: "c \ 0 \ b = a * c \ b / c = a" by (simp add: divide_inverse mult.assoc) lemma eq_divide_imp: "c \ 0 \ a * c = b \ a = b / c" by (drule sym) (simp add: divide_inverse mult.assoc) lemma add_divide_eq_iff [field_simps]: "z \ 0 \ x + y / z = (x * z + y) / z" by (simp add: add_divide_distrib nonzero_eq_divide_eq) lemma divide_add_eq_iff [field_simps]: "z \ 0 \ x / z + y = (x + y * z) / z" by (simp add: add_divide_distrib nonzero_eq_divide_eq) lemma diff_divide_eq_iff [field_simps]: "z \ 0 \ x - y / z = (x * z - y) / z" by (simp add: diff_divide_distrib nonzero_eq_divide_eq eq_diff_eq) lemma minus_divide_add_eq_iff [field_simps]: "z \ 0 \ - (x / z) + y = (- x + y * z) / z" by (simp add: add_divide_distrib diff_divide_eq_iff) lemma divide_diff_eq_iff [field_simps]: "z \ 0 \ x / z - y = (x - y * z) / z" by (simp add: field_simps) lemma minus_divide_diff_eq_iff [field_simps]: "z \ 0 \ - (x / z) - y = (- x - y * z) / z" by (simp add: divide_diff_eq_iff[symmetric]) lemma division_ring_divide_zero [simp]: "a / 0 = 0" by (simp add: divide_inverse) lemma divide_self_if [simp]: "a / a = (if a = 0 then 0 else 1)" by simp lemma inverse_nonzero_iff_nonzero [simp]: "inverse a = 0 \ a = 0" by (rule iffI) (fact inverse_zero_imp_zero, simp) lemma inverse_minus_eq [simp]: "inverse (- a) = - inverse a" proof cases assume "a=0" thus ?thesis by simp next assume "a\0" thus ?thesis by (simp add: nonzero_inverse_minus_eq) qed lemma inverse_inverse_eq [simp]: "inverse (inverse a) = a" proof cases assume "a=0" thus ?thesis by simp next assume "a\0" thus ?thesis by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_eq_imp_eq: "inverse a = inverse b \ a = b" by (drule arg_cong [where f="inverse"], simp) lemma inverse_eq_iff_eq [simp]: "inverse a = inverse b \ a = b" by (force dest!: inverse_eq_imp_eq) lemma mult_commute_imp_mult_inverse_commute: assumes "y * x = x * y" shows "inverse y * x = x * inverse y" proof (cases "y=0") case False hence "x * inverse y = inverse y * y * x * inverse y" by simp also have "\ = inverse y * (x * y * inverse y)" by (simp add: mult.assoc assms) finally show ?thesis by (simp add: mult.assoc False) qed simp lemmas mult_inverse_of_nat_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_nat_commute] lemma divide_divide_eq_left': "(a / b) / c = a / (c * b)" by (cases "b = 0 \ c = 0") (auto simp: divide_inverse mult.assoc nonzero_inverse_mult_distrib) lemma add_divide_eq_if_simps [field_split_simps, divide_simps]: "a + b / z = (if z = 0 then a else (a * z + b) / z)" "a / z + b = (if z = 0 then b else (a + b * z) / z)" "- (a / z) + b = (if z = 0 then b else (-a + b * z) / z)" "a - b / z = (if z = 0 then a else (a * z - b) / z)" "a / z - b = (if z = 0 then -b else (a - b * z) / z)" "- (a / z) - b = (if z = 0 then -b else (- a - b * z) / z)" by (simp_all add: add_divide_eq_iff divide_add_eq_iff diff_divide_eq_iff divide_diff_eq_iff minus_divide_diff_eq_iff) lemma [field_split_simps, divide_simps]: shows divide_eq_eq: "b / c = a \ (if c \ 0 then b = a * c else a = 0)" and eq_divide_eq: "a = b / c \ (if c \ 0 then a * c = b else a = 0)" and minus_divide_eq_eq: "- (b / c) = a \ (if c \ 0 then - b = a * c else a = 0)" and eq_minus_divide_eq: "a = - (b / c) \ (if c \ 0 then a * c = - b else a = 0)" by (auto simp add: field_simps) end subsection \Fields\ class field = comm_ring_1 + inverse + assumes field_inverse: "a \ 0 \ inverse a * a = 1" assumes field_divide_inverse: "a / b = a * inverse b" assumes field_inverse_zero: "inverse 0 = 0" begin subclass division_ring proof fix a :: 'a assume "a \ 0" thus "inverse a * a = 1" by (rule field_inverse) thus "a * inverse a = 1" by (simp only: mult.commute) next fix a b :: 'a show "a / b = a * inverse b" by (rule field_divide_inverse) next show "inverse 0 = 0" by (fact field_inverse_zero) qed subclass idom_divide proof fix b a assume "b \ 0" then show "a * b / b = a" by (simp add: divide_inverse ac_simps) next fix a show "a / 0 = 0" by (simp add: divide_inverse) qed text\There is no slick version using division by zero.\ lemma inverse_add: "a \ 0 \ b \ 0 \ inverse a + inverse b = (a + b) * inverse a * inverse b" by (simp add: division_ring_inverse_add ac_simps) lemma nonzero_mult_divide_mult_cancel_left [simp]: assumes [simp]: "c \ 0" shows "(c * a) / (c * b) = a / b" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "(c*a)/(c*b) = c * a * (inverse b * inverse c)" by (simp add: divide_inverse nonzero_inverse_mult_distrib) also have "... = a * inverse b * (inverse c * c)" by (simp only: ac_simps) also have "... = a * inverse b" by simp finally show ?thesis by (simp add: divide_inverse) qed lemma nonzero_mult_divide_mult_cancel_right [simp]: "c \ 0 \ (a * c) / (b * c) = a / b" using nonzero_mult_divide_mult_cancel_left [of c a b] by (simp add: ac_simps) lemma times_divide_eq_left [simp]: "(b / c) * a = (b * a) / c" by (simp add: divide_inverse ac_simps) lemma divide_inverse_commute: "a / b = inverse b * a" by (simp add: divide_inverse mult.commute) lemma add_frac_eq: assumes "y \ 0" and "z \ 0" shows "x / y + w / z = (x * z + w * y) / (y * z)" proof - have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)" using assms by simp also have "\ = (x * z + y * w) / (y * z)" by (simp only: add_divide_distrib) finally show ?thesis by (simp only: mult.commute) qed text\Special Cancellation Simprules for Division\ lemma nonzero_divide_mult_cancel_right [simp]: "b \ 0 \ b / (a * b) = 1 / a" using nonzero_mult_divide_mult_cancel_right [of b 1 a] by simp lemma nonzero_divide_mult_cancel_left [simp]: "a \ 0 \ a / (a * b) = 1 / b" using nonzero_mult_divide_mult_cancel_left [of a 1 b] by simp lemma nonzero_mult_divide_mult_cancel_left2 [simp]: "c \ 0 \ (c * a) / (b * c) = a / b" using nonzero_mult_divide_mult_cancel_left [of c a b] by (simp add: ac_simps) lemma nonzero_mult_divide_mult_cancel_right2 [simp]: "c \ 0 \ (a * c) / (c * b) = a / b" using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: ac_simps) lemma diff_frac_eq: "y \ 0 \ z \ 0 \ x / y - w / z = (x * z - w * y) / (y * z)" by (simp add: field_simps) lemma frac_eq_eq: "y \ 0 \ z \ 0 \ (x / y = w / z) = (x * z = w * y)" by (simp add: field_simps) lemma divide_minus1 [simp]: "x / - 1 = - x" using nonzero_minus_divide_right [of "1" x] by simp text\This version builds in division by zero while also re-orienting the right-hand side.\ lemma inverse_mult_distrib [simp]: "inverse (a * b) = inverse a * inverse b" proof cases assume "a \ 0 \ b \ 0" thus ?thesis by (simp add: nonzero_inverse_mult_distrib ac_simps) next assume "\ (a \ 0 \ b \ 0)" thus ?thesis by force qed lemma inverse_divide [simp]: "inverse (a / b) = b / a" by (simp add: divide_inverse mult.commute) text \Calculations with fractions\ text\There is a whole bunch of simp-rules just for class \field\ but none for class \field\ and \nonzero_divides\ because the latter are covered by a simproc.\ lemmas mult_divide_mult_cancel_left = nonzero_mult_divide_mult_cancel_left lemmas mult_divide_mult_cancel_right = nonzero_mult_divide_mult_cancel_right lemma divide_divide_eq_right [simp]: "a / (b / c) = (a * c) / b" by (simp add: divide_inverse ac_simps) lemma divide_divide_eq_left [simp]: "(a / b) / c = a / (b * c)" by (simp add: divide_inverse mult.assoc) lemma divide_divide_times_eq: "(x / y) / (z / w) = (x * w) / (y * z)" by simp text \Special Cancellation Simprules for Division\ lemma mult_divide_mult_cancel_left_if [simp]: shows "(c * a) / (c * b) = (if c = 0 then 0 else a / b)" by simp text \Division and Unary Minus\ lemma minus_divide_right: "- (a / b) = a / - b" by (simp add: divide_inverse) lemma divide_minus_right [simp]: "a / - b = - (a / b)" by (simp add: divide_inverse) lemma minus_divide_divide: "(- a) / (- b) = a / b" by (cases "b=0") (simp_all add: nonzero_minus_divide_divide) lemma inverse_eq_1_iff [simp]: "inverse x = 1 \ x = 1" using inverse_eq_iff_eq [of x 1] by simp lemma divide_eq_0_iff [simp]: "a / b = 0 \ a = 0 \ b = 0" by (simp add: divide_inverse) lemma divide_cancel_right [simp]: "a / c = b / c \ c = 0 \ a = b" by (cases "c=0") (simp_all add: divide_inverse) lemma divide_cancel_left [simp]: "c / a = c / b \ c = 0 \ a = b" by (cases "c=0") (simp_all add: divide_inverse) lemma divide_eq_1_iff [simp]: "a / b = 1 \ b \ 0 \ a = b" by (cases "b=0") (simp_all add: right_inverse_eq) lemma one_eq_divide_iff [simp]: "1 = a / b \ b \ 0 \ a = b" by (simp add: eq_commute [of 1]) lemma divide_eq_minus_1_iff: "(a / b = - 1) \ b \ 0 \ a = - b" using divide_eq_1_iff by fastforce lemma times_divide_times_eq: "(x / y) * (z / w) = (x * z) / (y * w)" by simp lemma add_frac_num: "y \ 0 \ x / y + z = (x + z * y) / y" by (simp add: add_divide_distrib) lemma add_num_frac: "y \ 0 \ z + x / y = (x + z * y) / y" by (simp add: add_divide_distrib add.commute) lemma dvd_field_iff: "a dvd b \ (a = 0 \ b = 0)" proof (cases "a = 0") case False then have "b = a * (b / a)" by (simp add: field_simps) then have "a dvd b" .. with False show ?thesis by simp qed simp lemma inj_divide_right [simp]: "inj (\b. b / a) \ a \ 0" proof - have "(\b. b / a) = (*) (inverse a)" by (simp add: field_simps fun_eq_iff) then have "inj (\y. y / a) \ inj ((*) (inverse a))" by simp also have "\ \ inverse a \ 0" by simp also have "\ \ a \ 0" by simp finally show ?thesis by simp qed end class field_char_0 = field + ring_char_0 subsection \Ordered fields\ class field_abs_sgn = field + idom_abs_sgn begin lemma sgn_inverse [simp]: "sgn (inverse a) = inverse (sgn a)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * inverse a = 1" by simp then have "sgn (a * inverse a) = sgn 1" by simp then have "sgn a * sgn (inverse a) = 1" by (simp add: sgn_mult) then have "inverse (sgn a) * (sgn a * sgn (inverse a)) = inverse (sgn a) * 1" by simp then have "(inverse (sgn a) * sgn a) * sgn (inverse a) = inverse (sgn a)" by (simp add: ac_simps) with False show ?thesis by (simp add: sgn_eq_0_iff) qed lemma abs_inverse [simp]: "\inverse a\ = inverse \a\" proof - from sgn_mult_abs [of "inverse a"] sgn_mult_abs [of a] have "inverse (sgn a) * \inverse a\ = inverse (sgn a * \a\)" by simp then show ?thesis by (auto simp add: sgn_eq_0_iff) qed lemma sgn_divide [simp]: "sgn (a / b) = sgn a / sgn b" unfolding divide_inverse sgn_mult by simp lemma abs_divide [simp]: "\a / b\ = \a\ / \b\" unfolding divide_inverse abs_mult by simp end class linordered_field = field + linordered_idom begin lemma positive_imp_inverse_positive: assumes a_gt_0: "0 < a" shows "0 < inverse a" proof - have "0 < a * inverse a" by (simp add: a_gt_0 [THEN less_imp_not_eq2]) thus "0 < inverse a" by (simp add: a_gt_0 [THEN less_not_sym] zero_less_mult_iff) qed lemma negative_imp_inverse_negative: "a < 0 \ inverse a < 0" using positive_imp_inverse_positive [of "-a"] by (simp add: nonzero_inverse_minus_eq less_imp_not_eq) lemma inverse_le_imp_le: assumes invle: "inverse a \ inverse b" and apos: "0 < a" shows "b \ a" proof (rule classical) assume "\ b \ a" hence "a < b" by (simp add: linorder_not_le) hence bpos: "0 < b" by (blast intro: apos less_trans) hence "a * inverse a \ a * inverse b" by (simp add: apos invle less_imp_le mult_left_mono) hence "(a * inverse a) * b \ (a * inverse b) * b" by (simp add: bpos less_imp_le mult_right_mono) thus "b \ a" by (simp add: mult.assoc apos bpos less_imp_not_eq2) qed lemma inverse_positive_imp_positive: assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" shows "0 < a" proof - have "0 < inverse (inverse a)" using inv_gt_0 by (rule positive_imp_inverse_positive) thus "0 < a" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_negative_imp_negative: assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" shows "a < 0" proof - have "inverse (inverse a) < 0" using inv_less_0 by (rule negative_imp_inverse_negative) thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma linordered_field_no_lb: "\x. \y. y < x" proof fix x::'a have m1: "- (1::'a) < 0" by simp from add_strict_right_mono[OF m1, where c=x] have "(- 1) + x < x" by simp thus "\y. y < x" by blast qed lemma linordered_field_no_ub: "\ x. \y. y > x" proof fix x::'a have m1: " (1::'a) > 0" by simp from add_strict_right_mono[OF m1, where c=x] have "1 + x > x" by simp thus "\y. y > x" by blast qed lemma less_imp_inverse_less: assumes less: "a < b" and apos: "0 < a" shows "inverse b < inverse a" proof (rule ccontr) assume "\ inverse b < inverse a" hence "inverse a \ inverse b" by simp hence "\ (a < b)" by (simp add: not_less inverse_le_imp_le [OF _ apos]) thus False by (rule notE [OF _ less]) qed lemma inverse_less_imp_less: assumes "inverse a < inverse b" "0 < a" shows "b < a" proof - have "a \ b" using assms by (simp add: less_le) moreover have "b \ a" using assms by (force simp: less_le dest: inverse_le_imp_le) ultimately show ?thesis by (simp add: less_le) qed text\Both premises are essential. Consider -1 and 1.\ lemma inverse_less_iff_less [simp]: "0 < a \ 0 < b \ inverse a < inverse b \ b < a" by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) lemma le_imp_inverse_le: "a \ b \ 0 < a \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less) lemma inverse_le_iff_le [simp]: "0 < a \ 0 < b \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) text\These results refer to both operands being negative. The opposite-sign case is trivial, since inverse preserves signs.\ lemma inverse_le_imp_le_neg: assumes "inverse a \ inverse b" "b < 0" shows "b \ a" proof (rule classical) assume "\ b \ a" with \b < 0\ have "a < 0" by force with assms show "b \ a" using inverse_le_imp_le [of "-b" "-a"] by (simp add: nonzero_inverse_minus_eq) qed lemma less_imp_inverse_less_neg: assumes "a < b" "b < 0" shows "inverse b < inverse a" proof - have "a < 0" using assms by (blast intro: less_trans) with less_imp_inverse_less [of "-b" "-a"] show ?thesis by (simp add: nonzero_inverse_minus_eq assms) qed lemma inverse_less_imp_less_neg: assumes "inverse a < inverse b" "b < 0" shows "b < a" proof (rule classical) assume "\ b < a" with \b < 0\ have "a < 0" by force with inverse_less_imp_less [of "-b" "-a"] show ?thesis by (simp add: nonzero_inverse_minus_eq assms) qed lemma inverse_less_iff_less_neg [simp]: "a < 0 \ b < 0 \ inverse a < inverse b \ b < a" using inverse_less_iff_less [of "-b" "-a"] by (simp del: inverse_less_iff_less add: nonzero_inverse_minus_eq) lemma le_imp_inverse_le_neg: "a \ b \ b < 0 \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less_neg) lemma inverse_le_iff_le_neg [simp]: "a < 0 \ b < 0 \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) lemma one_less_inverse: "0 < a \ a < 1 \ 1 < inverse a" using less_imp_inverse_less [of a 1, unfolded inverse_1] . lemma one_le_inverse: "0 < a \ a \ 1 \ 1 \ inverse a" using le_imp_inverse_le [of a 1, unfolded inverse_1] . lemma pos_le_divide_eq [field_simps]: assumes "0 < c" shows "a \ b / c \ a * c \ b" proof - from assms have "a \ b / c \ a * c \ (b / c) * c" using mult_le_cancel_right [of a c "b * inverse c"] by (auto simp add: field_simps) also have "... \ a * c \ b" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_less_divide_eq [field_simps]: assumes "0 < c" shows "a < b / c \ a * c < b" proof - from assms have "a < b / c \ a * c < (b / c) * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... = (a*c < b)" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_less_divide_eq [field_simps]: assumes "c < 0" shows "a < b / c \ b < a * c" proof - from assms have "a < b / c \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_le_divide_eq [field_simps]: assumes "c < 0" shows "a \ b / c \ b \ a * c" proof - from assms have "a \ b / c \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b * inverse c" c a] by (auto simp add: field_simps) also have "... \ b \ a * c" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_divide_le_eq [field_simps]: assumes "0 < c" shows "b / c \ a \ b \ a * c" proof - from assms have "b / c \ a \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b / c" c a] by auto also have "... \ b \ a * c" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_divide_less_eq [field_simps]: assumes "0 < c" shows "b / c < a \ b < a * c" proof - from assms have "b / c < a \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_divide_le_eq [field_simps]: assumes "c < 0" shows "b / c \ a \ a * c \ b" proof - from assms have "b / c \ a \ a * c \ (b / c) * c" using mult_le_cancel_right [of a c "b / c"] by auto also have "... \ a * c \ b" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_divide_less_eq [field_simps]: assumes "c < 0" shows "b / c < a \ a * c < b" proof - from assms have "b / c < a \ a * c < b / c * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... \ a * c < b" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed text\The following \field_simps\ rules are necessary, as minus is always moved atop of division but we want to get rid of division.\ lemma pos_le_minus_divide_eq [field_simps]: "0 < c \ a \ - (b / c) \ a * c \ - b" unfolding minus_divide_left by (rule pos_le_divide_eq) lemma neg_le_minus_divide_eq [field_simps]: "c < 0 \ a \ - (b / c) \ - b \ a * c" unfolding minus_divide_left by (rule neg_le_divide_eq) lemma pos_less_minus_divide_eq [field_simps]: "0 < c \ a < - (b / c) \ a * c < - b" unfolding minus_divide_left by (rule pos_less_divide_eq) lemma neg_less_minus_divide_eq [field_simps]: "c < 0 \ a < - (b / c) \ - b < a * c" unfolding minus_divide_left by (rule neg_less_divide_eq) lemma pos_minus_divide_less_eq [field_simps]: "0 < c \ - (b / c) < a \ - b < a * c" unfolding minus_divide_left by (rule pos_divide_less_eq) lemma neg_minus_divide_less_eq [field_simps]: "c < 0 \ - (b / c) < a \ a * c < - b" unfolding minus_divide_left by (rule neg_divide_less_eq) lemma pos_minus_divide_le_eq [field_simps]: "0 < c \ - (b / c) \ a \ - b \ a * c" unfolding minus_divide_left by (rule pos_divide_le_eq) lemma neg_minus_divide_le_eq [field_simps]: "c < 0 \ - (b / c) \ a \ a * c \ - b" unfolding minus_divide_left by (rule neg_divide_le_eq) lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) lemma divide_pos_pos[simp]: "0 < x \ 0 < y \ 0 < x / y" by(simp add:field_simps) lemma divide_nonneg_pos: "0 \ x \ 0 < y \ 0 \ x / y" by(simp add:field_simps) lemma divide_neg_pos: "x < 0 \ 0 < y \ x / y < 0" by(simp add:field_simps) lemma divide_nonpos_pos: "x \ 0 \ 0 < y \ x / y \ 0" by(simp add:field_simps) lemma divide_pos_neg: "0 < x \ y < 0 \ x / y < 0" by(simp add:field_simps) lemma divide_nonneg_neg: "0 \ x \ y < 0 \ x / y \ 0" by(simp add:field_simps) lemma divide_neg_neg: "x < 0 \ y < 0 \ 0 < x / y" by(simp add:field_simps) lemma divide_nonpos_neg: "x \ 0 \ y < 0 \ 0 \ x / y" by(simp add:field_simps) lemma divide_strict_right_mono: "\a < b; 0 < c\ \ a / c < b / c" by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono positive_imp_inverse_positive) lemma divide_strict_right_mono_neg: assumes "b < a" "c < 0" shows "a / c < b / c" proof - have "b / - c < a / - c" by (rule divide_strict_right_mono) (use assms in auto) then show ?thesis by (simp add: less_imp_not_eq) qed text\The last premise ensures that \<^term>\a\ and \<^term>\b\ have the same sign\ lemma divide_strict_left_mono: "\b < a; 0 < c; 0 < a*b\ \ c / a < c / b" by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono) lemma divide_left_mono: "\b \ a; 0 \ c; 0 < a*b\ \ c / a \ c / b" by (auto simp: field_simps zero_less_mult_iff mult_right_mono) lemma divide_strict_left_mono_neg: "\a < b; c < 0; 0 < a*b\ \ c / a < c / b" by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono_neg) lemma mult_imp_div_pos_le: "0 < y \ x \ z * y \ x / y \ z" by (subst pos_divide_le_eq, assumption+) lemma mult_imp_le_div_pos: "0 < y \ z * y \ x \ z \ x / y" by(simp add:field_simps) lemma mult_imp_div_pos_less: "0 < y \ x < z * y \ x / y < z" by(simp add:field_simps) lemma mult_imp_less_div_pos: "0 < y \ z * y < x \ z < x / y" by(simp add:field_simps) lemma frac_le: assumes "0 \ y" "x \ y" "0 < w" "w \ z" shows "x / z \ y / w" proof (rule mult_imp_div_pos_le) show "z > 0" using assms by simp have "x \ y * z / w" proof (rule mult_imp_le_div_pos [OF \0 < w\]) show "x * w \ y * z" using assms by (auto intro: mult_mono) qed also have "... = y / w * z" by simp finally show "x \ y / w * z" . qed lemma frac_less: assumes "0 \ x" "x < y" "0 < w" "w \ z" shows "x / z < y / w" proof (rule mult_imp_div_pos_less) show "z > 0" using assms by simp have "x < y * z / w" proof (rule mult_imp_less_div_pos [OF \0 < w\]) show "x * w < y * z" using assms by (auto intro: mult_less_le_imp_less) qed also have "... = y / w * z" by simp finally show "x < y / w * z" . qed lemma frac_less2: assumes "0 < x" "x \ y" "0 < w" "w < z" shows "x / z < y / w" proof (rule mult_imp_div_pos_less) show "z > 0" using assms by simp show "x < y / w * z" using assms by (force intro: mult_imp_less_div_pos mult_le_less_imp_less) qed lemma less_half_sum: "a < b \ a < (a+b) / (1+1)" by (simp add: field_simps zero_less_two) lemma gt_half_sum: "a < b \ (a+b)/(1+1) < b" by (simp add: field_simps zero_less_two) subclass unbounded_dense_linorder proof fix x y :: 'a from less_add_one show "\y. x < y" .. from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono) then have "x - 1 < x + 1 - 1" by simp then have "x - 1 < x" by (simp add: algebra_simps) then show "\y. y < x" .. show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) qed subclass field_abs_sgn .. lemma inverse_sgn [simp]: "inverse (sgn a) = sgn a" by (cases a 0 rule: linorder_cases) simp_all lemma divide_sgn [simp]: "a / sgn b = a * sgn b" by (cases b 0 rule: linorder_cases) simp_all lemma nonzero_abs_inverse: "a \ 0 \ \inverse a\ = inverse \a\" by (rule abs_inverse) lemma nonzero_abs_divide: "b \ 0 \ \a / b\ = \a\ / \b\" by (rule abs_divide) lemma field_le_epsilon: assumes e: "\e. 0 < e \ x \ y + e" shows "x \ y" proof (rule dense_le) fix t assume "t < x" hence "0 < x - t" by (simp add: less_diff_eq) from e [OF this] have "x + 0 \ x + (y - t)" by (simp add: algebra_simps) then have "0 \ y - t" by (simp only: add_le_cancel_left) then show "t \ y" by (simp add: algebra_simps) qed lemma inverse_positive_iff_positive [simp]: "(0 < inverse a) = (0 < a)" proof (cases "a = 0") case False then show ?thesis by (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive) qed auto lemma inverse_negative_iff_negative [simp]: "(inverse a < 0) = (a < 0)" proof (cases "a = 0") case False then show ?thesis by (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative) qed auto lemma inverse_nonnegative_iff_nonnegative [simp]: "0 \ inverse a \ 0 \ a" by (simp add: not_less [symmetric]) lemma inverse_nonpositive_iff_nonpositive [simp]: "inverse a \ 0 \ a \ 0" by (simp add: not_less [symmetric]) lemma one_less_inverse_iff: "1 < inverse x \ 0 < x \ x < 1" using less_trans[of 1 x 0 for x] by (cases x 0 rule: linorder_cases) (auto simp add: field_simps) lemma one_le_inverse_iff: "1 \ inverse x \ 0 < x \ x \ 1" proof (cases "x = 1") case True then show ?thesis by simp next case False then have "inverse x \ 1" by simp then have "1 \ inverse x" by blast then have "1 \ inverse x \ 1 < inverse x" by (simp add: le_less) with False show ?thesis by (auto simp add: one_less_inverse_iff) qed lemma inverse_less_1_iff: "inverse x < 1 \ x \ 0 \ 1 < x" by (simp add: not_le [symmetric] one_le_inverse_iff) lemma inverse_le_1_iff: "inverse x \ 1 \ x \ 0 \ 1 \ x" by (simp add: not_less [symmetric] one_less_inverse_iff) lemma [field_split_simps, divide_simps]: shows le_divide_eq: "a \ b / c \ (if 0 < c then a * c \ b else if c < 0 then b \ a * c else a \ 0)" and divide_le_eq: "b / c \ a \ (if 0 < c then b \ a * c else if c < 0 then a * c \ b else 0 \ a)" and less_divide_eq: "a < b / c \ (if 0 < c then a * c < b else if c < 0 then b < a * c else a < 0)" and divide_less_eq: "b / c < a \ (if 0 < c then b < a * c else if c < 0 then a * c < b else 0 < a)" and le_minus_divide_eq: "a \ - (b / c) \ (if 0 < c then a * c \ - b else if c < 0 then - b \ a * c else a \ 0)" and minus_divide_le_eq: "- (b / c) \ a \ (if 0 < c then - b \ a * c else if c < 0 then a * c \ - b else 0 \ a)" and less_minus_divide_eq: "a < - (b / c) \ (if 0 < c then a * c < - b else if c < 0 then - b < a * c else a < 0)" and minus_divide_less_eq: "- (b / c) < a \ (if 0 < c then - b < a * c else if c < 0 then a * c < - b else 0 < a)" by (auto simp: field_simps not_less dest: order.antisym) text \Division and Signs\ lemma shows zero_less_divide_iff: "0 < a / b \ 0 < a \ 0 < b \ a < 0 \ b < 0" and divide_less_0_iff: "a / b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" and zero_le_divide_iff: "0 \ a / b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" and divide_le_0_iff: "a / b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" by (auto simp add: field_split_simps) text \Division and the Number One\ text\Simplify expressions equated with 1\ lemma zero_eq_1_divide_iff [simp]: "0 = 1 / a \ a = 0" by (cases "a = 0") (auto simp: field_simps) lemma one_divide_eq_0_iff [simp]: "1 / a = 0 \ a = 0" using zero_eq_1_divide_iff[of a] by simp text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ lemma zero_le_divide_1_iff [simp]: "0 \ 1 / a \ 0 \ a" by (simp add: zero_le_divide_iff) lemma zero_less_divide_1_iff [simp]: "0 < 1 / a \ 0 < a" by (simp add: zero_less_divide_iff) lemma divide_le_0_1_iff [simp]: "1 / a \ 0 \ a \ 0" by (simp add: divide_le_0_iff) lemma divide_less_0_1_iff [simp]: "1 / a < 0 \ a < 0" by (simp add: divide_less_0_iff) lemma divide_right_mono: "\a \ b; 0 \ c\ \ a/c \ b/c" by (force simp add: divide_strict_right_mono le_less) lemma divide_right_mono_neg: "a \ b \ c \ 0 \ b / c \ a / c" by (auto dest: divide_right_mono [of _ _ "- c"]) lemma divide_left_mono_neg: "a \ b \ c \ 0 \ 0 < a * b \ c / a \ c / b" by (auto simp add: mult.commute dest: divide_left_mono [of _ _ "- c"]) lemma inverse_le_iff: "inverse a \ inverse b \ (0 < a * b \ b \ a) \ (a * b \ 0 \ a \ b)" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: field_simps zero_less_mult_iff mult_le_0_iff) lemma inverse_less_iff: "inverse a < inverse b \ (0 < a * b \ b < a) \ (a * b \ 0 \ a < b)" by (subst less_le) (auto simp: inverse_le_iff) lemma divide_le_cancel: "a / c \ b / c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: divide_inverse mult_le_cancel_right) lemma divide_less_cancel: "a / c < b / c \ (0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0" by (auto simp add: divide_inverse mult_less_cancel_right) text\Simplify quotients that are compared with the value 1.\ lemma le_divide_eq_1: "(1 \ b / a) = ((0 < a \ a \ b) \ (a < 0 \ b \ a))" by (auto simp add: le_divide_eq) lemma divide_le_eq_1: "(b / a \ 1) = ((0 < a \ b \ a) \ (a < 0 \ a \ b) \ a=0)" by (auto simp add: divide_le_eq) lemma less_divide_eq_1: "(1 < b / a) = ((0 < a \ a < b) \ (a < 0 \ b < a))" by (auto simp add: less_divide_eq) lemma divide_less_eq_1: "(b / a < 1) = ((0 < a \ b < a) \ (a < 0 \ a < b) \ a=0)" by (auto simp add: divide_less_eq) lemma divide_nonneg_nonneg [simp]: "0 \ x \ 0 \ y \ 0 \ x / y" by (auto simp add: field_split_simps) lemma divide_nonpos_nonpos: "x \ 0 \ y \ 0 \ 0 \ x / y" by (auto simp add: field_split_simps) lemma divide_nonneg_nonpos: "0 \ x \ y \ 0 \ x / y \ 0" by (auto simp add: field_split_simps) lemma divide_nonpos_nonneg: "x \ 0 \ 0 \ y \ x / y \ 0" by (auto simp add: field_split_simps) text \Conditional Simplification Rules: No Case Splits\ lemma le_divide_eq_1_pos [simp]: "0 < a \ (1 \ b/a) = (a \ b)" by (auto simp add: le_divide_eq) lemma le_divide_eq_1_neg [simp]: "a < 0 \ (1 \ b/a) = (b \ a)" by (auto simp add: le_divide_eq) lemma divide_le_eq_1_pos [simp]: "0 < a \ (b/a \ 1) = (b \ a)" by (auto simp add: divide_le_eq) lemma divide_le_eq_1_neg [simp]: "a < 0 \ (b/a \ 1) = (a \ b)" by (auto simp add: divide_le_eq) lemma less_divide_eq_1_pos [simp]: "0 < a \ (1 < b/a) = (a < b)" by (auto simp add: less_divide_eq) lemma less_divide_eq_1_neg [simp]: "a < 0 \ (1 < b/a) = (b < a)" by (auto simp add: less_divide_eq) lemma divide_less_eq_1_pos [simp]: "0 < a \ (b/a < 1) = (b < a)" by (auto simp add: divide_less_eq) lemma divide_less_eq_1_neg [simp]: "a < 0 \ b/a < 1 \ a < b" by (auto simp add: divide_less_eq) lemma eq_divide_eq_1 [simp]: "(1 = b/a) = ((a \ 0 \ a = b))" by (auto simp add: eq_divide_eq) lemma divide_eq_eq_1 [simp]: "(b/a = 1) = ((a \ 0 \ a = b))" by (auto simp add: divide_eq_eq) lemma abs_div_pos: "0 < y \ \x\ / y = \x / y\" by (simp add: order_less_imp_le) lemma zero_le_divide_abs_iff [simp]: "(0 \ a / \b\) = (0 \ a \ b = 0)" by (auto simp: zero_le_divide_iff) lemma divide_le_0_abs_iff [simp]: "(a / \b\ \ 0) = (a \ 0 \ b = 0)" by (auto simp: divide_le_0_iff) lemma field_le_mult_one_interval: assumes *: "\z. \ 0 < z ; z < 1 \ \ z * x \ y" shows "x \ y" proof (cases "0 < x") assume "0 < x" thus ?thesis using dense_le_bounded[of 0 1 "y/x"] * unfolding le_divide_eq if_P[OF \0 < x\] by simp next assume "\0 < x" hence "x \ 0" by simp obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1::'a"] by auto hence "x \ s * x" using mult_le_cancel_right[of 1 x s] \x \ 0\ by auto also note *[OF s] finally show ?thesis . qed text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: assumes "u \ v" "0 \ r" "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms using divide_le_eq_1 by fastforce moreover have "0 \ v - u" using assms by simp ultimately have "(r/s) * (v - u) \ 1 * (v - u)" by (rule mult_right_mono) then show ?thesis by (simp add: field_simps) qed end text \Min/max Simplification Rules\ lemma min_mult_distrib_left: fixes x::"'a::linordered_idom" shows "p * min x y = (if 0 \ p then min (p*x) (p*y) else max (p*x) (p*y))" by (auto simp add: min_def max_def mult_le_cancel_left) lemma min_mult_distrib_right: fixes x::"'a::linordered_idom" shows "min x y * p = (if 0 \ p then min (x*p) (y*p) else max (x*p) (y*p))" by (auto simp add: min_def max_def mult_le_cancel_right) lemma min_divide_distrib_right: fixes x::"'a::linordered_field" shows "min x y / p = (if 0 \ p then min (x/p) (y/p) else max (x/p) (y/p))" by (simp add: min_mult_distrib_right divide_inverse) lemma max_mult_distrib_left: fixes x::"'a::linordered_idom" shows "p * max x y = (if 0 \ p then max (p*x) (p*y) else min (p*x) (p*y))" by (auto simp add: min_def max_def mult_le_cancel_left) lemma max_mult_distrib_right: fixes x::"'a::linordered_idom" shows "max x y * p = (if 0 \ p then max (x*p) (y*p) else min (x*p) (y*p))" by (auto simp add: min_def max_def mult_le_cancel_right) lemma max_divide_distrib_right: fixes x::"'a::linordered_field" shows "max x y / p = (if 0 \ p then max (x/p) (y/p) else min (x/p) (y/p))" by (simp add: max_mult_distrib_right divide_inverse) hide_fact (open) field_inverse field_divide_inverse field_inverse_zero code_identifier code_module Fields \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Int.thy b/src/HOL/Int.thy --- a/src/HOL/Int.thy +++ b/src/HOL/Int.thy @@ -1,2272 +1,2262 @@ (* Title: HOL/Int.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) section \The Integers as Equivalence Classes over Pairs of Natural Numbers\ theory Int imports Quotient Groups_Big Fun_Def begin subsection \Definition of integers as a quotient type\ definition intrel :: "(nat \ nat) \ (nat \ nat) \ bool" where "intrel = (\(x, y) (u, v). x + v = u + y)" lemma intrel_iff [simp]: "intrel (x, y) (u, v) \ x + v = u + y" by (simp add: intrel_def) quotient_type int = "nat \ nat" / "intrel" morphisms Rep_Integ Abs_Integ proof (rule equivpI) show "reflp intrel" by (auto simp: reflp_def) show "symp intrel" by (auto simp: symp_def) show "transp intrel" by (auto simp: transp_def) qed lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: "(\x y. z = Abs_Integ (x, y) \ P) \ P" by (induct z) auto subsection \Integers form a commutative ring\ instantiation int :: comm_ring_1 begin lift_definition zero_int :: "int" is "(0, 0)" . lift_definition one_int :: "int" is "(1, 0)" . lift_definition plus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + u, y + v)" by clarsimp lift_definition uminus_int :: "int \ int" is "\(x, y). (y, x)" by clarsimp lift_definition minus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + v, y + u)" by clarsimp lift_definition times_int :: "int \ int \ int" is "\(x, y) (u, v). (x*u + y*v, x*v + y*u)" proof (unfold intrel_def, clarify) fix s t u v w x y z :: nat assume "s + v = u + t" and "w + z = y + x" then have "(s + v) * w + (u + t) * x + u * (w + z) + v * (y + x) = (u + t) * w + (s + v) * x + u * (y + x) + v * (w + z)" by simp then show "(s * w + t * x) + (u * z + v * y) = (u * y + v * z) + (s * x + t * w)" by (simp add: algebra_simps) qed instance by standard (transfer; clarsimp simp: algebra_simps)+ end abbreviation int :: "nat \ int" where "int \ of_nat" lemma int_def: "int n = Abs_Integ (n, 0)" by (induct n) (simp add: zero_int.abs_eq, simp add: one_int.abs_eq plus_int.abs_eq) lemma int_transfer [transfer_rule]: includes lifting_syntax shows "rel_fun (=) pcr_int (\n. (n, 0)) int" by (simp add: rel_fun_def int.pcr_cr_eq cr_int_def int_def) lemma int_diff_cases: obtains (diff) m n where "z = int m - int n" by transfer clarsimp subsection \Integers are totally ordered\ instantiation int :: linorder begin lift_definition less_eq_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v \ u + y" by auto lift_definition less_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v < u + y" by auto instance by standard (transfer, force)+ end instantiation int :: distrib_lattice begin definition "(inf :: int \ int \ int) = min" definition "(sup :: int \ int \ int) = max" instance by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2) end subsection \Ordering properties of arithmetic operations\ instance int :: ordered_cancel_ab_semigroup_add proof fix i j k :: int show "i \ j \ k + i \ k + j" by transfer clarsimp qed text \Strict Monotonicity of Multiplication.\ text \Strict, in 1st argument; proof is by induction on \k > 0\.\ lemma zmult_zless_mono2_lemma: "i < j \ 0 < k \ int k * i < int k * j" for i j :: int proof (induct k) case 0 then show ?case by simp next case (Suc k) then show ?case by (cases "k = 0") (simp_all add: distrib_right add_strict_mono) qed lemma zero_le_imp_eq_int: assumes "k \ (0::int)" shows "\n. k = int n" proof - have "b \ a \ \n::nat. a = n + b" for a b using exI[of _ "a - b"] by simp with assms show ?thesis by transfer auto qed lemma zero_less_imp_eq_int: assumes "k > (0::int)" shows "\n>0. k = int n" proof - have "b < a \ \n::nat. n>0 \ a = n + b" for a b using exI[of _ "a - b"] by simp with assms show ?thesis by transfer auto qed lemma zmult_zless_mono2: "i < j \ 0 < k \ k * i < k * j" for i j k :: int by (drule zero_less_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma) text \The integers form an ordered integral domain.\ instantiation int :: linordered_idom begin definition zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)" instance proof fix i j k :: int show "i < j \ 0 < k \ k * i < k * j" by (rule zmult_zless_mono2) show "\i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) show "sgn (i::int) = (if i=0 then 0 else if 0 w + 1 \ z" for w z :: int by transfer clarsimp lemma zless_iff_Suc_zadd: "w < z \ (\n. z = w + int (Suc n))" for w z :: int proof - have "\a b c d. a + d < c + b \ \n. c + b = Suc (a + n + d)" proof - fix a b c d :: nat assume "a + d < c + b" then have "c + b = Suc (a + (c + b - Suc (a + d)) + d) " by arith then show "\n. c + b = Suc (a + n + d)" by (rule exI) qed then show ?thesis by transfer auto qed lemma zabs_less_one_iff [simp]: "\z\ < 1 \ z = 0" (is "?lhs \ ?rhs") for z :: int proof assume ?rhs then show ?lhs by simp next assume ?lhs with zless_imp_add1_zle [of "\z\" 1] have "\z\ + 1 \ 1" by simp then have "\z\ \ 0" by simp then show ?rhs by simp qed subsection \Embedding of the Integers into any \ring_1\: \of_int\\ context ring_1 begin lift_definition of_int :: "int \ 'a" is "\(i, j). of_nat i - of_nat j" by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_0 [simp]: "of_int 0 = 0" by transfer simp lemma of_int_1 [simp]: "of_int 1 = 1" by transfer simp lemma of_int_add [simp]: "of_int (w + z) = of_int w + of_int z" by transfer (clarsimp simp add: algebra_simps) lemma of_int_minus [simp]: "of_int (- z) = - (of_int z)" by (transfer fixing: uminus) clarsimp lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" using of_int_add [of w "- z"] by simp lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" by (transfer fixing: times) (clarsimp simp add: algebra_simps) lemma mult_of_int_commute: "of_int x * y = y * of_int x" by (transfer fixing: times) (auto simp: algebra_simps mult_of_nat_commute) text \Collapse nested embeddings.\ lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n" by (induct n) auto lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k" by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric]) lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k" by simp lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n" by (induct n) simp_all lemma of_int_of_bool [simp]: "of_int (of_bool P) = of_bool P" by auto end context ring_char_0 begin lemma of_int_eq_iff [simp]: "of_int w = of_int z \ w = z" by transfer (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) text \Special cases where either operand is zero.\ lemma of_int_eq_0_iff [simp]: "of_int z = 0 \ z = 0" using of_int_eq_iff [of z 0] by simp lemma of_int_0_eq_iff [simp]: "0 = of_int z \ z = 0" using of_int_eq_iff [of 0 z] by simp lemma of_int_eq_1_iff [iff]: "of_int z = 1 \ z = 1" using of_int_eq_iff [of z 1] by simp lemma numeral_power_eq_of_int_cancel_iff [simp]: "numeral x ^ n = of_int y \ numeral x ^ n = y" using of_int_eq_iff[of "numeral x ^ n" y, unfolded of_int_numeral of_int_power] . lemma of_int_eq_numeral_power_cancel_iff [simp]: "of_int y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_int_cancel_iff [of x n y] by (metis (mono_tags)) lemma neg_numeral_power_eq_of_int_cancel_iff [simp]: "(- numeral x) ^ n = of_int y \ (- numeral x) ^ n = y" using of_int_eq_iff[of "(- numeral x) ^ n" y] by simp lemma of_int_eq_neg_numeral_power_cancel_iff [simp]: "of_int y = (- numeral x) ^ n \ y = (- numeral x) ^ n" using neg_numeral_power_eq_of_int_cancel_iff[of x n y] by (metis (mono_tags)) lemma of_int_eq_of_int_power_cancel_iff[simp]: "(of_int b) ^ w = of_int x \ b ^ w = x" by (metis of_int_power of_int_eq_iff) lemma of_int_power_eq_of_int_cancel_iff[simp]: "of_int x = (of_int b) ^ w \ x = b ^ w" by (metis of_int_eq_of_int_power_cancel_iff) end context linordered_idom begin text \Every \linordered_idom\ has characteristic zero.\ subclass ring_char_0 .. lemma of_int_le_iff [simp]: "of_int w \ of_int z \ w \ z" by (transfer fixing: less_eq) (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_less_iff [simp]: "of_int w < of_int z \ w < z" by (simp add: less_le order_less_le) lemma of_int_0_le_iff [simp]: "0 \ of_int z \ 0 \ z" using of_int_le_iff [of 0 z] by simp lemma of_int_le_0_iff [simp]: "of_int z \ 0 \ z \ 0" using of_int_le_iff [of z 0] by simp lemma of_int_0_less_iff [simp]: "0 < of_int z \ 0 < z" using of_int_less_iff [of 0 z] by simp lemma of_int_less_0_iff [simp]: "of_int z < 0 \ z < 0" using of_int_less_iff [of z 0] by simp lemma of_int_1_le_iff [simp]: "1 \ of_int z \ 1 \ z" using of_int_le_iff [of 1 z] by simp lemma of_int_le_1_iff [simp]: "of_int z \ 1 \ z \ 1" using of_int_le_iff [of z 1] by simp lemma of_int_1_less_iff [simp]: "1 < of_int z \ 1 < z" using of_int_less_iff [of 1 z] by simp lemma of_int_less_1_iff [simp]: "of_int z < 1 \ z < 1" using of_int_less_iff [of z 1] by simp lemma of_int_pos: "z > 0 \ of_int z > 0" by simp lemma of_int_nonneg: "z \ 0 \ of_int z \ 0" by simp lemma of_int_abs [simp]: "of_int \x\ = \of_int x\" by (auto simp add: abs_if) lemma of_int_lessD: assumes "\of_int n\ < x" shows "n = 0 \ x > 1" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 < x" using assms by (rule le_less_trans) then show ?thesis .. qed lemma of_int_leD: assumes "\of_int n\ \ x" shows "n = 0 \ 1 \ x" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 \ x" using assms by (rule order_trans) then show ?thesis .. qed lemma numeral_power_le_of_int_cancel_iff [simp]: "numeral x ^ n \ of_int a \ numeral x ^ n \ a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_le_iff) lemma of_int_le_numeral_power_cancel_iff [simp]: "of_int a \ numeral x ^ n \ a \ numeral x ^ n" by (metis (mono_tags) local.numeral_power_eq_of_int_cancel_iff of_int_le_iff) lemma numeral_power_less_of_int_cancel_iff [simp]: "numeral x ^ n < of_int a \ numeral x ^ n < a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma of_int_less_numeral_power_cancel_iff [simp]: "of_int a < numeral x ^ n \ a < numeral x ^ n" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma neg_numeral_power_le_of_int_cancel_iff [simp]: "(- numeral x) ^ n \ of_int a \ (- numeral x) ^ n \ a" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma of_int_le_neg_numeral_power_cancel_iff [simp]: "of_int a \ (- numeral x) ^ n \ a \ (- numeral x) ^ n" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma neg_numeral_power_less_of_int_cancel_iff [simp]: "(- numeral x) ^ n < of_int a \ (- numeral x) ^ n < a" using of_int_less_iff[of "(- numeral x) ^ n" a] by simp lemma of_int_less_neg_numeral_power_cancel_iff [simp]: "of_int a < (- numeral x) ^ n \ a < (- numeral x::int) ^ n" using of_int_less_iff[of a "(- numeral x) ^ n"] by simp lemma of_int_le_of_int_power_cancel_iff[simp]: "(of_int b) ^ w \ of_int x \ b ^ w \ x" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_power_le_of_int_cancel_iff[simp]: "of_int x \ (of_int b) ^ w\ x \ b ^ w" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_less_of_int_power_cancel_iff[simp]: "(of_int b) ^ w < of_int x \ b ^ w < x" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_power_less_of_int_cancel_iff[simp]: "of_int x < (of_int b) ^ w\ x < b ^ w" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_max: "of_int (max x y) = max (of_int x) (of_int y)" by (auto simp: max_def) lemma of_int_min: "of_int (min x y) = min (of_int x) (of_int y)" by (auto simp: min_def) end context division_ring begin lemmas mult_inverse_of_int_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_int_commute] end text \Comparisons involving \<^term>\of_int\.\ lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \ z = numeral n" using of_int_eq_iff by fastforce lemma of_int_le_numeral_iff [simp]: "of_int z \ (numeral n :: 'a::linordered_idom) \ z \ numeral n" using of_int_le_iff [of z "numeral n"] by simp lemma of_int_numeral_le_iff [simp]: "(numeral n :: 'a::linordered_idom) \ of_int z \ numeral n \ z" using of_int_le_iff [of "numeral n"] by simp lemma of_int_less_numeral_iff [simp]: "of_int z < (numeral n :: 'a::linordered_idom) \ z < numeral n" using of_int_less_iff [of z "numeral n"] by simp lemma of_int_numeral_less_iff [simp]: "(numeral n :: 'a::linordered_idom) < of_int z \ numeral n < z" using of_int_less_iff [of "numeral n" z] by simp lemma of_nat_less_of_int_iff: "(of_nat n::'a::linordered_idom) < of_int x \ int n < x" by (metis of_int_of_nat_eq of_int_less_iff) lemma of_int_eq_id [simp]: "of_int = id" proof show "of_int z = id z" for z by (cases z rule: int_diff_cases) simp qed instance int :: no_top proof fix x::int have "x < x + 1" by simp then show "\y. x < y" by (rule exI) qed instance int :: no_bot proof fix x::int have "x - 1< x" by simp then show "\y. y < x" by (rule exI) qed subsection \Magnitude of an Integer, as a Natural Number: \nat\\ lift_definition nat :: "int \ nat" is "\(x, y). x - y" by auto lemma nat_int [simp]: "nat (int n) = n" by transfer simp lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" by transfer clarsimp lemma nat_0_le: "0 \ z \ int (nat z) = z" by simp lemma nat_le_0 [simp]: "z \ 0 \ nat z = 0" by transfer clarsimp lemma nat_le_eq_zle: "0 < w \ 0 \ z \ nat w \ nat z \ w \ z" by transfer (clarsimp, arith) text \An alternative condition is \<^term>\0 \ w\.\ lemma nat_mono_iff: "0 < z \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma nat_less_eq_zless: "0 \ w \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma zless_nat_conj [simp]: "nat w < nat z \ 0 < z \ w < z" by transfer (clarsimp, arith) lemma nonneg_int_cases: assumes "0 \ k" obtains n where "k = int n" proof - from assms have "k = int (nat k)" by simp then show thesis by (rule that) qed lemma pos_int_cases: assumes "0 < k" obtains n where "k = int n" and "n > 0" proof - from assms have "0 \ k" by simp then obtain n where "k = int n" by (rule nonneg_int_cases) moreover have "n > 0" using \k = int n\ assms by simp ultimately show thesis by (rule that) qed lemma nonpos_int_cases: assumes "k \ 0" obtains n where "k = - int n" proof - from assms have "- k \ 0" by simp then obtain n where "- k = int n" by (rule nonneg_int_cases) then have "k = - int n" by simp then show thesis by (rule that) qed lemma neg_int_cases: assumes "k < 0" obtains n where "k = - int n" and "n > 0" proof - from assms have "- k > 0" by simp then obtain n where "- k = int n" and "- k > 0" by (blast elim: pos_int_cases) then have "k = - int n" and "n > 0" by simp_all then show thesis by (rule that) qed lemma nat_eq_iff: "nat w = m \ (if 0 \ w then w = int m else m = 0)" by transfer (clarsimp simp add: le_imp_diff_is_add) lemma nat_eq_iff2: "m = nat w \ (if 0 \ w then w = int m else m = 0)" using nat_eq_iff [of w m] by auto lemma nat_0 [simp]: "nat 0 = 0" by (simp add: nat_eq_iff) lemma nat_1 [simp]: "nat 1 = Suc 0" by (simp add: nat_eq_iff) lemma nat_numeral [simp]: "nat (numeral k) = numeral k" by (simp add: nat_eq_iff) lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0" by simp lemma nat_2: "nat 2 = Suc (Suc 0)" by simp lemma nat_less_iff: "0 \ w \ nat w < m \ w < of_nat m" by transfer (clarsimp, arith) lemma nat_le_iff: "nat x \ n \ x \ int n" by transfer (clarsimp simp add: le_diff_conv) lemma nat_mono: "x \ y \ nat x \ nat y" by transfer auto lemma nat_0_iff[simp]: "nat i = 0 \ i \ 0" for i :: int by transfer clarsimp lemma int_eq_iff: "of_nat m = z \ m = nat z \ 0 \ z" by (auto simp add: nat_eq_iff2) lemma zero_less_nat_eq [simp]: "0 < nat z \ 0 < z" using zless_nat_conj [of 0] by auto lemma nat_add_distrib: "0 \ z \ 0 \ z' \ nat (z + z') = nat z + nat z'" by transfer clarsimp lemma nat_diff_distrib': "0 \ x \ 0 \ y \ nat (x - y) = nat x - nat y" by transfer clarsimp lemma nat_diff_distrib: "0 \ z' \ z' \ z \ nat (z - z') = nat z - nat z'" by (rule nat_diff_distrib') auto lemma nat_zminus_int [simp]: "nat (- int n) = 0" by transfer simp lemma le_nat_iff: "k \ 0 \ n \ nat k \ int n \ k" by transfer auto lemma zless_nat_eq_int_zless: "m < nat z \ int m < z" by transfer (clarsimp simp add: less_diff_conv) lemma (in ring_1) of_nat_nat [simp]: "0 \ z \ of_nat (nat z) = of_int z" by transfer (clarsimp simp add: of_nat_diff) lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')" by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral) lemma nat_abs_triangle_ineq: "nat \k + l\ \ nat \k\ + nat \l\" by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq) lemma nat_of_bool [simp]: "nat (of_bool P) = of_bool P" by auto lemma split_nat [linarith_split]: "P (nat i) \ ((\n. i = int n \ P n) \ (i < 0 \ P 0))" (is "?P = (?L \ ?R)") for i :: int proof (cases "i < 0") case True then show ?thesis by auto next case False have "?P = ?L" proof assume ?P then show ?L using False by auto next assume ?L moreover from False have "int (nat i) = i" by (simp add: not_less) ultimately show ?P by simp qed with False show ?thesis by simp qed lemma all_nat: "(\x. P x) \ (\x\0. P (nat x))" by (auto split: split_nat) lemma ex_nat: "(\x. P x) \ (\x. 0 \ x \ P (nat x))" proof assume "\x. P x" then obtain x where "P x" .. then have "int x \ 0 \ P (nat (int x))" by simp then show "\x\0. P (nat x)" .. next assume "\x\0. P (nat x)" then show "\x. P x" by auto qed text \For termination proofs:\ lemma measure_function_int[measure_function]: "is_measure (nat \ abs)" .. subsection \Lemmas about the Function \<^term>\of_nat\ and Orderings\ lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)" by (simp add: order_less_le del: of_nat_Suc) lemma negative_zless [iff]: "- (int (Suc n)) < int m" by (rule negative_zless_0 [THEN order_less_le_trans], simp) lemma negative_zle_0: "- int n \ 0" by (simp add: minus_le_iff) lemma negative_zle [iff]: "- int n \ int m" by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) lemma not_zle_0_negative [simp]: "\ 0 \ - int (Suc n)" by (subst le_minus_iff) (simp del: of_nat_Suc) lemma int_zle_neg: "int n \ - int m \ n = 0 \ m = 0" by transfer simp lemma not_int_zless_negative [simp]: "\ int n < - int m" by (simp add: linorder_not_less) lemma negative_eq_positive [simp]: "- int n = of_nat m \ n = 0 \ m = 0" by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) lemma zle_iff_zadd: "w \ z \ (\n. z = w + int n)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by auto next assume ?lhs then have "0 \ z - w" by simp then obtain n where "z - w = int n" using zero_le_imp_eq_int [of "z - w"] by blast then have "z = w + int n" by simp then show ?rhs .. qed lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z" by simp -text \ - This version is proved for all ordered rings, not just integers! - It is proved here because attribute \linarith_split\ is not available - in theory \Rings\. - But is it really better than just rewriting with \abs_if\? -\ -lemma abs_split [linarith_split, no_atp]: "P \a\ \ (0 \ a \ P a) \ (a < 0 \ P (- a))" - for a :: "'a::linordered_idom" - by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) - lemma negD: assumes "x < 0" shows "\n. x = - (int (Suc n))" proof - have "\a b. a < b \ \n. Suc (a + n) = b" proof - fix a b:: nat assume "a < b" then have "Suc (a + (b - Suc a)) = b" by arith then show "\n. Suc (a + n) = b" by (rule exI) qed with assms show ?thesis by transfer auto qed subsection \Cases and induction\ text \ Now we replace the case analysis rule by a more conventional one: whether an integer is negative or not. \ text \This version is symmetric in the two subgoals.\ lemma int_cases2 [case_names nonneg nonpos, cases type: int]: "(\n. z = int n \ P) \ (\n. z = - (int n) \ P) \ P" by (cases "z < 0") (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym]) text \This is the default, with a negative case.\ lemma int_cases [case_names nonneg neg, cases type: int]: assumes pos: "\n. z = int n \ P" and neg: "\n. z = - (int (Suc n)) \ P" shows P proof (cases "z < 0") case True with neg show ?thesis by (blast dest!: negD) next case False with pos show ?thesis by (force simp add: linorder_not_less dest: nat_0_le [THEN sym]) qed lemma int_cases3 [case_names zero pos neg]: fixes k :: int assumes "k = 0 \ P" and "\n. k = int n \ n > 0 \ P" and "\n. k = - int n \ n > 0 \ P" shows "P" proof (cases k "0::int" rule: linorder_cases) case equal with assms(1) show P by simp next case greater then have *: "nat k > 0" by simp moreover from * have "k = int (nat k)" by auto ultimately show P using assms(2) by blast next case less then have *: "nat (- k) > 0" by simp moreover from * have "k = - int (nat (- k))" by auto ultimately show P using assms(3) by blast qed lemma int_of_nat_induct [case_names nonneg neg, induct type: int]: "(\n. P (int n)) \ (\n. P (- (int (Suc n)))) \ P z" by (cases z) auto lemma sgn_mult_dvd_iff [simp]: "sgn r * l dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int by (cases r rule: int_cases3) auto lemma mult_sgn_dvd_iff [simp]: "l * sgn r dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps) lemma dvd_sgn_mult_iff [simp]: "l dvd sgn r * k \ l dvd k \ r = 0" for k l r :: int by (cases r rule: int_cases3) simp_all lemma dvd_mult_sgn_iff [simp]: "l dvd k * sgn r \ l dvd k \ r = 0" for k l r :: int using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps) lemma int_sgnE: fixes k :: int obtains n and l where "k = sgn l * int n" proof - have "k = sgn k * int (nat \k\)" by (simp add: sgn_mult_abs) then show ?thesis .. qed subsubsection \Binary comparisons\ text \Preliminaries\ lemma le_imp_0_less: fixes z :: int assumes le: "0 \ z" shows "0 < 1 + z" proof - have "0 \ z" by fact also have "\ < z + 1" by (rule less_add_one) also have "\ = 1 + z" by (simp add: ac_simps) finally show "0 < 1 + z" . qed lemma odd_less_0_iff: "1 + z + z < 0 \ z < 0" for z :: int proof (cases z) case (nonneg n) then show ?thesis by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le]) next case (neg n) then show ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1 add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric]) qed subsubsection \Comparisons, for Ordered Rings\ lemma odd_nonzero: "1 + z + z \ 0" for z :: int proof (cases z) case (nonneg n) have le: "0 \ z + z" by (simp add: nonneg add_increasing) then show ?thesis using le_imp_0_less [OF le] by (auto simp: ac_simps) next case (neg n) show ?thesis proof assume eq: "1 + z + z = 0" have "0 < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) also have "\ = - (1 + z + z)" by (simp add: neg add.assoc [symmetric]) also have "\ = 0" by (simp add: eq) finally have "0<0" .. then show False by blast qed qed subsection \The Set of Integers\ context ring_1 begin definition Ints :: "'a set" ("\") where "\ = range of_int" lemma Ints_of_int [simp]: "of_int z \ \" by (simp add: Ints_def) lemma Ints_of_nat [simp]: "of_nat n \ \" using Ints_of_int [of "of_nat n"] by simp lemma Ints_0 [simp]: "0 \ \" using Ints_of_int [of "0"] by simp lemma Ints_1 [simp]: "1 \ \" using Ints_of_int [of "1"] by simp lemma Ints_numeral [simp]: "numeral n \ \" by (subst of_nat_numeral [symmetric], rule Ints_of_nat) lemma Ints_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (force simp add: Ints_def simp flip: of_int_add intro: range_eqI) lemma Ints_minus [simp]: "a \ \ \ -a \ \" by (force simp add: Ints_def simp flip: of_int_minus intro: range_eqI) lemma minus_in_Ints_iff: "-x \ \ \ x \ \" using Ints_minus[of x] Ints_minus[of "-x"] by auto lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (force simp add: Ints_def simp flip: of_int_diff intro: range_eqI) lemma Ints_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (force simp add: Ints_def simp flip: of_int_mult intro: range_eqI) lemma Ints_power [simp]: "a \ \ \ a ^ n \ \" by (induct n) simp_all lemma Ints_cases [cases set: Ints]: assumes "q \ \" obtains (of_int) z where "q = of_int z" unfolding Ints_def proof - from \q \ \\ have "q \ range of_int" unfolding Ints_def . then obtain z where "q = of_int z" .. then show thesis .. qed lemma Ints_induct [case_names of_int, induct set: Ints]: "q \ \ \ (\z. P (of_int z)) \ P q" by (rule Ints_cases) auto lemma Nats_subset_Ints: "\ \ \" unfolding Nats_def Ints_def by (rule subsetI, elim imageE, hypsubst, subst of_int_of_nat_eq[symmetric], rule imageI) simp_all lemma Nats_altdef1: "\ = {of_int n |n. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {of_int n |n. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp next fix x :: 'a assume "x \ \" then obtain n where "x = of_nat n" by (auto elim!: Nats_cases) then have "x = of_int (int n)" by simp also have "int n \ 0" by simp then have "of_int (int n) \ {of_int n |n. n \ 0}" by blast finally show "x \ {of_int n |n. n \ 0}" . qed end lemma Ints_sum [intro]: "(\x. x \ A \ f x \ \) \ sum f A \ \" by (induction A rule: infinite_finite_induct) auto lemma Ints_prod [intro]: "(\x. x \ A \ f x \ \) \ prod f A \ \" by (induction A rule: infinite_finite_induct) auto lemma (in linordered_idom) Ints_abs [simp]: shows "a \ \ \ abs a \ \" by (auto simp: abs_if) lemma (in linordered_idom) Nats_altdef2: "\ = {n \ \. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {n \ \. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp qed (auto elim!: Nats_cases) lemma (in idom_divide) of_int_divide_in_Ints: "of_int a div of_int b \ \" if "b dvd a" proof - from that obtain c where "a = b * c" .. then show ?thesis by (cases "of_int b = 0") simp_all qed text \The premise involving \<^term>\Ints\ prevents \<^term>\a = 1/2\.\ lemma Ints_double_eq_0_iff: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "a + a = 0 \ a = 0" (is "?lhs \ ?rhs") proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume ?rhs then show ?lhs by simp next assume ?lhs with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp then have "z + z = 0" by (simp only: of_int_eq_iff) then have "z = 0" by (simp only: double_zero) with a show ?rhs by simp qed qed lemma Ints_odd_nonzero: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "1 + a + a \ 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume "1 + a + a = 0" with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp then have "1 + z + z = 0" by (simp only: of_int_eq_iff) with odd_nonzero show False by blast qed qed lemma Nats_numeral [simp]: "numeral w \ \" using of_nat_in_Nats [of "numeral w"] by simp lemma Ints_odd_less_0: fixes a :: "'a::linordered_idom" assumes in_Ints: "a \ \" shows "1 + a + a < 0 \ a < 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. with a have "1 + a + a < 0 \ of_int (1 + z + z) < (of_int 0 :: 'a)" by simp also have "\ \ z < 0" by (simp only: of_int_less_iff odd_less_0_iff) also have "\ \ a < 0" by (simp add: a) finally show ?thesis . qed subsection \\<^term>\sum\ and \<^term>\prod\\ context semiring_1 begin lemma of_nat_sum [simp]: "of_nat (sum f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context ring_1 begin lemma of_int_sum [simp]: "of_int (sum f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_semiring_1 begin lemma of_nat_prod [simp]: "of_nat (prod f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_ring_1 begin lemma of_int_prod [simp]: "of_int (prod f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end subsection \Setting up simplification procedures\ ML_file \Tools/int_arith.ML\ declaration \K ( Lin_Arith.add_discrete_type \<^type_name>\Int.int\ #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle} #> Lin_Arith.add_inj_thms @{thms of_nat_le_iff [THEN iffD2] of_nat_eq_iff [THEN iffD2]} #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ int\) #> Lin_Arith.add_simps @{thms of_int_0 of_int_1 of_int_add of_int_mult of_int_numeral of_int_neg_numeral nat_0 nat_1 diff_nat_numeral nat_numeral neg_less_iff_less True_implies_equals distrib_left [where a = "numeral v" for v] distrib_left [where a = "- numeral v" for v] div_by_1 div_0 times_divide_eq_right times_divide_eq_left minus_divide_left [THEN sym] minus_divide_right [THEN sym] add_divide_distrib diff_divide_distrib of_int_minus of_int_diff of_int_of_nat_eq} #> Lin_Arith.add_simprocs [Int_Arith.zero_one_idom_simproc] )\ simproc_setup fast_arith ("(m::'a::linordered_idom) < n" | "(m::'a::linordered_idom) \ n" | "(m::'a::linordered_idom) = n") = \K Lin_Arith.simproc\ subsection\More Inequality Reasoning\ lemma zless_add1_eq: "w < z + 1 \ w < z \ w = z" for w z :: int by arith lemma add1_zle_eq: "w + 1 \ z \ w < z" for w z :: int by arith lemma zle_diff1_eq [simp]: "w \ z - 1 \ w < z" for w z :: int by arith lemma zle_add1_eq_le [simp]: "w < z + 1 \ w \ z" for w z :: int by arith lemma int_one_le_iff_zero_less: "1 \ z \ 0 < z" for z :: int by arith lemma Ints_nonzero_abs_ge1: fixes x:: "'a :: linordered_idom" assumes "x \ Ints" "x \ 0" shows "1 \ abs x" proof (rule Ints_cases [OF \x \ Ints\]) fix z::int assume "x = of_int z" with \x \ 0\ show "1 \ \x\" apply (auto simp: abs_if) by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq) qed lemma Ints_nonzero_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; abs x < 1\ \ x = 0" using Ints_nonzero_abs_ge1 [of x] by auto lemma Ints_eq_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; y \ Ints\ \ x = y \ abs (x-y) < 1" using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1) subsection \The functions \<^term>\nat\ and \<^term>\int\\ text \Simplify the term \<^term>\w + - z\.\ lemma one_less_nat_eq [simp]: "Suc 0 < nat z \ 1 < z" using zless_nat_conj [of 1 z] by auto lemma int_eq_iff_numeral [simp]: "int m = numeral v \ m = numeral v" by (simp add: int_eq_iff) lemma nat_abs_int_diff: "nat \int a - int b\ = (if a \ b then b - a else a - b)" by auto lemma nat_int_add: "nat (int a + int b) = a + b" by auto context ring_1 begin lemma of_int_of_nat [nitpick_simp]: "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" proof (cases "k < 0") case True then have "0 \ - k" by simp then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) with True show ?thesis by simp next case False then show ?thesis by (simp add: not_less) qed end lemma transfer_rule_of_int: includes lifting_syntax fixes R :: "'a::ring_1 \ 'b::ring_1 \ bool" assumes [transfer_rule]: "R 0 0" "R 1 1" "(R ===> R ===> R) (+) (+)" "(R ===> R) uminus uminus" shows "((=) ===> R) of_int of_int" proof - note assms note transfer_rule_of_nat [transfer_rule] have [transfer_rule]: "((=) ===> R) of_nat of_nat" by transfer_prover show ?thesis by (unfold of_int_of_nat [abs_def]) transfer_prover qed lemma nat_mult_distrib: fixes z z' :: int assumes "0 \ z" shows "nat (z * z') = nat z * nat z'" proof (cases "0 \ z'") case False with assms have "z * z' \ 0" by (simp add: not_le mult_le_0_iff) then have "nat (z * z') = 0" by simp moreover from False have "nat z' = 0" by simp ultimately show ?thesis by simp next case True with assms have ge_0: "z * z' \ 0" by (simp add: zero_le_mult_iff) show ?thesis by (rule injD [of "of_nat :: nat \ int", OF inj_of_nat]) (simp only: of_nat_mult of_nat_nat [OF True] of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp) qed lemma nat_mult_distrib_neg: assumes "z \ (0::int)" shows "nat (z * z') = nat (- z) * nat (- z')" (is "?L = ?R") proof - have "?L = nat (- z * - z')" using assms by auto also have "... = ?R" by (rule nat_mult_distrib) (use assms in auto) finally show ?thesis . qed lemma nat_abs_mult_distrib: "nat \w * z\ = nat \w\ * nat \z\" by (cases "z = 0 \ w = 0") (auto simp add: abs_if nat_mult_distrib [symmetric] nat_mult_distrib_neg [symmetric] mult_less_0_iff) lemma int_in_range_abs [simp]: "int n \ range abs" proof (rule range_eqI) show "int n = \int n\" by simp qed lemma range_abs_Nats [simp]: "range abs = (\ :: int set)" proof - have "\k\ \ \" for k :: int by (cases k) simp_all moreover have "k \ range abs" if "k \ \" for k :: int using that by induct simp ultimately show ?thesis by blast qed lemma Suc_nat_eq_nat_zadd1: "0 \ z \ Suc (nat z) = nat (1 + z)" for z :: int by (rule sym) (simp add: nat_eq_iff) lemma diff_nat_eq_if: "nat z - nat z' = (if z' < 0 then nat z else let d = z - z' in if d < 0 then 0 else nat d)" by (simp add: Let_def nat_diff_distrib [symmetric]) lemma nat_numeral_diff_1 [simp]: "numeral v - (1::nat) = nat (numeral v - 1)" using diff_nat_numeral [of v Num.One] by simp subsection \Induction principles for int\ text \Well-founded segments of the integers.\ definition int_ge_less_than :: "int \ (int \ int) set" where "int_ge_less_than d = {(z', z). d \ z' \ z' < z}" lemma wf_int_ge_less_than: "wf (int_ge_less_than d)" proof - have "int_ge_less_than d \ measure (\z. nat (z - d))" by (auto simp add: int_ge_less_than_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed text \ This variant looks odd, but is typical of the relations suggested by RankFinder.\ definition int_ge_less_than2 :: "int \ (int \ int) set" where "int_ge_less_than2 d = {(z',z). d \ z \ z' < z}" lemma wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" proof - have "int_ge_less_than2 d \ measure (\z. nat (1 + z - d))" by (auto simp add: int_ge_less_than2_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed (* `set:int': dummy construction *) theorem int_ge_induct [case_names base step, induct set: int]: fixes i :: int assumes ge: "k \ i" and base: "P k" and step: "\i. k \ i \ P i \ P (i + 1)" shows "P i" proof - have "\i::int. n = nat (i - k) \ k \ i \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat ((i - 1) - k)" by arith moreover have k: "k \ i - 1" using Suc.prems by arith ultimately have "P (i - 1)" by (rule Suc.hyps) from step [OF k this] show ?case by simp qed with ge show ?thesis by fast qed (* `set:int': dummy construction *) theorem int_gr_induct [case_names base step, induct set: int]: fixes i k :: int assumes "k < i" "P (k + 1)" "\i. k < i \ P i \ P (i + 1)" shows "P i" proof - have "k+1 \ i" using assms by auto then show ?thesis by (induction i rule: int_ge_induct) (auto simp: assms) qed theorem int_le_induct [consumes 1, case_names base step]: fixes i k :: int assumes le: "i \ k" and base: "P k" and step: "\i. i \ k \ P i \ P (i - 1)" shows "P i" proof - have "\i::int. n = nat(k-i) \ i \ k \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat (k - (i + 1))" by arith moreover have k: "i + 1 \ k" using Suc.prems by arith ultimately have "P (i + 1)" by (rule Suc.hyps) from step[OF k this] show ?case by simp qed with le show ?thesis by fast qed theorem int_less_induct [consumes 1, case_names base step]: fixes i k :: int assumes "i < k" "P (k - 1)" "\i. i < k \ P i \ P (i - 1)" shows "P i" proof - have "i \ k-1" using assms by auto then show ?thesis by (induction i rule: int_le_induct) (auto simp: assms) qed theorem int_induct [case_names base step1 step2]: fixes k :: int assumes base: "P k" and step1: "\i. k \ i \ P i \ P (i + 1)" and step2: "\i. k \ i \ P i \ P (i - 1)" shows "P i" proof - have "i \ k \ i \ k" by arith then show ?thesis proof assume "i \ k" then show ?thesis using base by (rule int_ge_induct) (fact step1) next assume "i \ k" then show ?thesis using base by (rule int_le_induct) (fact step2) qed qed subsection \Intermediate value theorems\ lemma nat_ivt_aux: "\\if (Suc i) - f i\ \ 1; f 0 \ k; k \ f n\ \ \i \ n. f i = k" for m n :: nat and k :: int proof (induct n) case (Suc n) show ?case proof (cases "k = f (Suc n)") case False with Suc have "k \ f n" by auto with Suc show ?thesis by (auto simp add: abs_if split: if_split_asm intro: le_SucI) qed (use Suc in auto) qed auto lemma nat_intermed_int_val: fixes m n :: nat and k :: int assumes "\i. m \ i \ i < n \ \f (Suc i) - f i\ \ 1" "m \ n" "f m \ k" "k \ f n" shows "\i. m \ i \ i \ n \ f i = k" proof - obtain i where "i \ n - m" "k = f (m + i)" using nat_ivt_aux [of "n - m" "f \ plus m" k] assms by auto with assms show ?thesis using exI[of _ "m + i"] by auto qed lemma nat0_intermed_int_val: "\i\n. f i = k" if "\if (i + 1) - f i\ \ 1" "f 0 \ k" "k \ f n" for n :: nat and k :: int using nat_intermed_int_val [of 0 n f k] that by auto subsection \Products and 1, by T. M. Rasmussen\ lemma abs_zmult_eq_1: fixes m n :: int assumes mn: "\m * n\ = 1" shows "\m\ = 1" proof - from mn have 0: "m \ 0" "n \ 0" by auto have "\ 2 \ \m\" proof assume "2 \ \m\" then have "2 * \n\ \ \m\ * \n\" by (simp add: mult_mono 0) also have "\ = \m * n\" by (simp add: abs_mult) also from mn have "\ = 1" by simp finally have "2 * \n\ \ 1" . with 0 show "False" by arith qed with 0 show ?thesis by auto qed lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \ m = 1 \ m = - 1" for m n :: int using abs_zmult_eq_1 [of m n] by arith lemma pos_zmult_eq_1_iff: fixes m n :: int assumes "0 < m" shows "m * n = 1 \ m = 1 \ n = 1" proof - from assms have "m * n = 1 \ m = 1" by (auto dest: pos_zmult_eq_1_iff_lemma) then show ?thesis by (auto dest: pos_zmult_eq_1_iff_lemma) qed lemma zmult_eq_1_iff: "m * n = 1 \ (m = 1 \ n = 1) \ (m = - 1 \ n = - 1)" (is "?L = ?R") for m n :: int proof assume L: ?L show ?R using pos_zmult_eq_1_iff_lemma [OF L] L by force qed auto lemma infinite_UNIV_int [simp]: "\ finite (UNIV::int set)" proof assume "finite (UNIV::int set)" moreover have "inj (\i::int. 2 * i)" by (rule injI) simp ultimately have "surj (\i::int. 2 * i)" by (rule finite_UNIV_inj_surj) then obtain i :: int where "1 = 2 * i" by (rule surjE) then show False by (simp add: pos_zmult_eq_1_iff) qed subsection \The divides relation\ lemma zdvd_antisym_nonneg: "0 \ m \ 0 \ n \ m dvd n \ n dvd m \ m = n" for m n :: int by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff) lemma zdvd_antisym_abs: fixes a b :: int assumes "a dvd b" and "b dvd a" shows "\a\ = \b\" proof (cases "a = 0") case True with assms show ?thesis by simp next case False from \a dvd b\ obtain k where k: "b = a * k" unfolding dvd_def by blast from \b dvd a\ obtain k' where k': "a = b * k'" unfolding dvd_def by blast from k k' have "a = a * k * k'" by simp with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1" using \a \ 0\ by (simp add: mult.assoc) then have "k = 1 \ k' = 1 \ k = -1 \ k' = -1" by (simp add: zmult_eq_1_iff) with k k' show ?thesis by auto qed lemma zdvd_zdiffD: "k dvd m - n \ k dvd n \ k dvd m" for k m n :: int using dvd_add_right_iff [of k "- n" m] by simp lemma zdvd_reduce: "k dvd n + k * m \ k dvd n" for k m n :: int using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps) lemma dvd_imp_le_int: fixes d i :: int assumes "i \ 0" and "d dvd i" shows "\d\ \ \i\" proof - from \d dvd i\ obtain k where "i = d * k" .. with \i \ 0\ have "k \ 0" by auto then have "1 \ \k\" and "0 \ \d\" by auto then have "\d\ * 1 \ \d\ * \k\" by (rule mult_left_mono) with \i = d * k\ show ?thesis by (simp add: abs_mult) qed lemma zdvd_not_zless: fixes m n :: int assumes "0 < m" and "m < n" shows "\ n dvd m" proof from assms have "0 < n" by auto assume "n dvd m" then obtain k where k: "m = n * k" .. with \0 < m\ have "0 < n * k" by auto with \0 < n\ have "0 < k" by (simp add: zero_less_mult_iff) with k \0 < n\ \m < n\ have "n * k < n * 1" by simp with \0 < n\ \0 < k\ show False unfolding mult_less_cancel_left by auto qed lemma zdvd_mult_cancel: fixes k m n :: int assumes d: "k * m dvd k * n" and "k \ 0" shows "m dvd n" proof - from d obtain h where h: "k * n = k * m * h" unfolding dvd_def by blast have "n = m * h" proof (rule ccontr) assume "\ ?thesis" with \k \ 0\ have "k * n \ k * (m * h)" by simp with h show False by (simp add: mult.assoc) qed then show ?thesis by simp qed lemma int_dvd_int_iff [simp]: "int m dvd int n \ m dvd n" proof - have "m dvd n" if "int n = int m * k" for k proof (cases k) case (nonneg q) with that have "n = m * q" by (simp del: of_nat_mult add: of_nat_mult [symmetric]) then show ?thesis .. next case (neg q) with that have "int n = int m * (- int (Suc q))" by simp also have "\ = - (int m * int (Suc q))" by (simp only: mult_minus_right) also have "\ = - int (m * Suc q)" by (simp only: of_nat_mult [symmetric]) finally have "- int (m * Suc q) = int n" .. then show ?thesis by (simp only: negative_eq_positive) auto qed then show ?thesis by (auto simp add: dvd_def) qed lemma dvd_nat_abs_iff [simp]: "n dvd nat \k\ \ int n dvd k" proof - have "n dvd nat \k\ \ int n dvd int (nat \k\)" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma nat_abs_dvd_iff [simp]: "nat \k\ dvd n \ k dvd int n" proof - have "nat \k\ dvd n \ int (nat \k\) dvd int n" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma zdvd1_eq [simp]: "x dvd 1 \ \x\ = 1" (is "?lhs \ ?rhs") for x :: int proof assume ?lhs then have "nat \x\ dvd nat \1\" by (simp only: nat_abs_dvd_iff) simp then have "nat \x\ = 1" by simp then show ?rhs by (cases "x < 0") simp_all next assume ?rhs then have "x = 1 \ x = - 1" by auto then show ?lhs by (auto intro: dvdI) qed lemma zdvd_mult_cancel1: fixes m :: int assumes mp: "m \ 0" shows "m * n dvd m \ \n\ = 1" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (cases "n > 0") (auto simp add: minus_equation_iff) next assume ?lhs then have "m * n dvd m * 1" by simp from zdvd_mult_cancel[OF this mp] show ?rhs by (simp only: zdvd1_eq) qed lemma nat_dvd_iff: "nat z dvd m \ (if 0 \ z then z dvd int m else m = 0)" using nat_abs_dvd_iff [of z m] by (cases "z \ 0") auto lemma eq_nat_nat_iff: "0 \ z \ 0 \ z' \ nat z = nat z' \ z = z'" by (auto elim: nonneg_int_cases) lemma nat_power_eq: "0 \ z \ nat (z ^ n) = nat z ^ n" by (induct n) (simp_all add: nat_mult_distrib) lemma numeral_power_eq_nat_cancel_iff [simp]: "numeral x ^ n = nat y \ numeral x ^ n = y" using nat_eq_iff2 by auto lemma nat_eq_numeral_power_cancel_iff [simp]: "nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_nat_cancel_iff[of x n y] by (metis (mono_tags)) lemma numeral_power_le_nat_cancel_iff [simp]: "numeral x ^ n \ nat a \ numeral x ^ n \ a" using nat_le_eq_zle[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_le_numeral_power_cancel_iff [simp]: "nat a \ numeral x ^ n \ a \ numeral x ^ n" by (simp add: nat_le_iff) lemma numeral_power_less_nat_cancel_iff [simp]: "numeral x ^ n < nat a \ numeral x ^ n < a" using nat_less_eq_zless[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_less_numeral_power_cancel_iff [simp]: "nat a < numeral x ^ n \ a < numeral x ^ n" using nat_less_eq_zless[of a "numeral x ^ n"] by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0]) lemma zdvd_imp_le: "z \ n" if "z dvd n" "0 < n" for n z :: int proof (cases n) case (nonneg n) show ?thesis by (cases z) (use nonneg dvd_imp_le that in auto) qed (use that in auto) lemma zdvd_period: fixes a d :: int assumes "a dvd d" shows "a dvd (x + t) \ a dvd ((x + c * d) + t)" (is "?lhs \ ?rhs") proof - from assms have "a dvd (x + t) \ a dvd ((x + t) + c * d)" by (simp add: dvd_add_left_iff) then show ?thesis by (simp add: ac_simps) qed subsection \Powers with integer exponents\ text \ The following allows writing powers with an integer exponent. While the type signature is very generic, most theorems will assume that the underlying type is a division ring or a field. The notation `powi' is inspired by the `powr' notation for real/complex exponentiation. \ definition power_int :: "'a :: {inverse, power} \ int \ 'a" (infixr "powi" 80) where "power_int x n = (if n \ 0 then x ^ nat n else inverse x ^ (nat (-n)))" lemma power_int_0_right [simp]: "power_int x 0 = 1" and power_int_1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) 1 = y" and power_int_minus1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) (-1) = inverse y" by (simp_all add: power_int_def) lemma power_int_of_nat [simp]: "power_int x (int n) = x ^ n" by (simp add: power_int_def) lemma power_int_numeral [simp]: "power_int x (numeral n) = x ^ numeral n" by (simp add: power_int_def) lemma int_cases4 [case_names nonneg neg]: fixes m :: int obtains n where "m = int n" | n where "n > 0" "m = -int n" proof (cases "m \ 0") case True thus ?thesis using that(1)[of "nat m"] by auto next case False thus ?thesis using that(2)[of "nat (-m)"] by auto qed context assumes "SORT_CONSTRAINT('a::division_ring)" begin lemma power_int_minus: "power_int (x::'a) (-n) = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_eq_0_iff [simp]: "power_int (x::'a) n = 0 \ x = 0 \ n \ 0" by (auto simp: power_int_def) lemma power_int_0_left_If: "power_int (0 :: 'a) m = (if m = 0 then 1 else 0)" by (auto simp: power_int_def) lemma power_int_0_left [simp]: "m \ 0 \ power_int (0 :: 'a) m = 0" by (simp add: power_int_0_left_If) lemma power_int_1_left [simp]: "power_int 1 n = (1 :: 'a :: division_ring)" by (auto simp: power_int_def) lemma power_diff_conv_inverse: "x \ 0 \ m \ n \ (x :: 'a) ^ (n - m) = x ^ n * inverse x ^ m" by (simp add: field_simps flip: power_add) lemma power_mult_inverse_distrib: "x ^ m * inverse (x :: 'a) = inverse x * x ^ m" proof (cases "x = 0") case [simp]: False show ?thesis proof (cases m) case (Suc m') have "x ^ Suc m' * inverse x = x ^ m'" by (subst power_Suc2) (auto simp: mult.assoc) also have "\ = inverse x * x ^ Suc m'" by (subst power_Suc) (auto simp: mult.assoc [symmetric]) finally show ?thesis using Suc by simp qed auto qed auto lemma power_mult_power_inverse_commute: "x ^ m * inverse (x :: 'a) ^ n = inverse x ^ n * x ^ m" proof (induction n) case (Suc n) have "x ^ m * inverse x ^ Suc n = (x ^ m * inverse x ^ n) * inverse x" by (simp only: power_Suc2 mult.assoc) also have "x ^ m * inverse x ^ n = inverse x ^ n * x ^ m" by (rule Suc) also have "\ * inverse x = (inverse x ^ n * inverse x) * x ^ m" by (simp add: mult.assoc power_mult_inverse_distrib) also have "\ = inverse x ^ (Suc n) * x ^ m" by (simp only: power_Suc2) finally show ?case . qed auto lemma power_int_add: assumes "x \ 0 \ m + n \ 0" shows "power_int (x::'a) (m + n) = power_int x m * power_int x n" proof (cases "x = 0") case True thus ?thesis using assms by (auto simp: power_int_0_left_If) next case [simp]: False show ?thesis proof (cases m n rule: int_cases4[case_product int_cases4]) case (nonneg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib power_add) next case (nonneg_neg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_neg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib add.commute simp flip: power_add) qed qed lemma power_int_add_1: assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = power_int x m * x" using assms by (subst power_int_add) auto lemma power_int_add_1': assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = x * power_int x m" using assms by (subst add.commute, subst power_int_add) auto lemma power_int_commutes: "power_int (x :: 'a) n * x = x * power_int x n" by (cases "x = 0") (auto simp flip: power_int_add_1 power_int_add_1') lemma power_int_inverse [field_simps, field_split_simps, divide_simps]: "power_int (inverse (x :: 'a)) n = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_mult: "power_int (x :: 'a) (m * n) = power_int (power_int x m) n" by (auto simp: power_int_def zero_le_mult_iff simp flip: power_mult power_inverse nat_mult_distrib) end context assumes "SORT_CONSTRAINT('a::field)" begin lemma power_int_diff: assumes "x \ 0 \ m \ n" shows "power_int (x::'a) (m - n) = power_int x m / power_int x n" using power_int_add[of x m "-n"] assms by (auto simp: field_simps power_int_minus) lemma power_int_minus_mult: "x \ 0 \ n \ 0 \ power_int (x :: 'a) (n - 1) * x = power_int x n" by (auto simp flip: power_int_add_1) lemma power_int_mult_distrib: "power_int (x * y :: 'a) m = power_int x m * power_int y m" by (auto simp: power_int_def power_mult_distrib) lemmas power_int_mult_distrib_numeral1 = power_int_mult_distrib [where x = "numeral w" for w, simp] lemmas power_int_mult_distrib_numeral2 = power_int_mult_distrib [where y = "numeral w" for w, simp] lemma power_int_divide_distrib: "power_int (x / y :: 'a) m = power_int x m / power_int y m" using power_int_mult_distrib[of x "inverse y" m] unfolding power_int_inverse by (simp add: field_simps) end lemma power_int_add_numeral [simp]: "power_int x (numeral m) * power_int x (numeral n) = power_int x (numeral (m + n))" for x :: "'a :: division_ring" by (simp add: power_int_add [symmetric]) lemma power_int_add_numeral2 [simp]: "power_int x (numeral m) * (power_int x (numeral n) * b) = power_int x (numeral (m + n)) * b" for x :: "'a :: division_ring" by (simp add: mult.assoc [symmetric]) lemma power_int_mult_numeral [simp]: "power_int (power_int x (numeral m)) (numeral n) = power_int x (numeral (m * n))" for x :: "'a :: division_ring" by (simp only: numeral_mult power_int_mult) lemma power_int_not_zero: "(x :: 'a :: division_ring) \ 0 \ n = 0 \ power_int x n \ 0" by (subst power_int_eq_0_iff) auto lemma power_int_one_over [field_simps, field_split_simps, divide_simps]: "power_int (1 / x :: 'a :: division_ring) n = 1 / power_int x n" using power_int_inverse[of x] by (simp add: divide_inverse) context assumes "SORT_CONSTRAINT('a :: linordered_field)" begin lemma power_int_numeral_neg_numeral [simp]: "power_int (numeral m) (-numeral n) = (inverse (numeral (Num.pow m n)) :: 'a)" by (simp add: power_int_minus) lemma zero_less_power_int [simp]: "0 < (x :: 'a) \ 0 < power_int x n" by (auto simp: power_int_def) lemma zero_le_power_int [simp]: "0 \ (x :: 'a) \ 0 \ power_int x n" by (auto simp: power_int_def) lemma power_int_mono: "(x :: 'a) \ y \ n \ 0 \ 0 \ x \ power_int x n \ power_int y n" by (cases n rule: int_cases4) (auto intro: power_mono) lemma one_le_power_int [simp]: "1 \ (x :: 'a) \ n \ 0 \ 1 \ power_int x n" using power_int_mono [of 1 x n] by simp lemma power_int_le_one: "0 \ (x :: 'a) \ n \ 0 \ x \ 1 \ power_int x n \ 1" using power_int_mono [of x 1 n] by simp lemma power_int_le_imp_le_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m \ power_int x n" "n \ 0" shows "m \ n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m \ x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_le_imp_less_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m < power_int x n" "n \ 0" shows "m < n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m < x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_strict_mono: "(a :: 'a :: linordered_field) < b \ 0 \ a \ 0 < n \ power_int a n < power_int b n" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_mono_iff [simp]: fixes a b :: "'a :: linordered_field" shows "\a \ 0; b \ 0; n > 0\ \ power_int a n \ power_int b n \ a \ b" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_strict_increasing: fixes a :: "'a :: linordered_field" assumes "n < N" "1 < a" shows "power_int a N > power_int a n" proof - have *: "a ^ nat (N - n) > a ^ 0" using assms by (intro power_strict_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ > power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_increasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "a \ 1" shows "power_int a N \ power_int a n" proof - have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_strict_decreasing: fixes a :: "'a :: linordered_field" assumes "n < N" "0 < a" "a < 1" shows "power_int a N < power_int a n" proof - have *: "a ^ nat (N - n) < a ^ 0" using assms by (intro power_strict_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ < power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_decreasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "0 \ a" "a \ 1" "a \ 0 \ N \ 0 \ n = 0" shows "power_int a N \ power_int a n" proof (cases "a = 0") case False have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms False by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed (use assms in \auto simp: power_int_0_left_If\) lemma one_less_power_int: "1 < (a :: 'a) \ 0 < n \ 1 < power_int a n" using power_int_strict_increasing[of 0 n a] by simp lemma power_int_abs: "\power_int a n :: 'a\ = power_int \a\ n" by (auto simp: power_int_def power_abs) lemma power_int_sgn [simp]: "sgn (power_int a n :: 'a) = power_int (sgn a) n" by (auto simp: power_int_def) lemma abs_power_int_minus [simp]: "\power_int (- a) n :: 'a\ = \power_int a n\" by (simp add: power_int_abs) lemma power_int_strict_antimono: assumes "(a :: 'a :: linordered_field) < b" "0 < a" "n < 0" shows "power_int a n > power_int b n" proof - have "inverse (power_int a (-n)) > inverse (power_int b (-n))" using assms by (intro less_imp_inverse_less power_int_strict_mono zero_less_power_int) auto thus ?thesis by (simp add: power_int_minus) qed lemma power_int_antimono: assumes "(a :: 'a :: linordered_field) \ b" "0 < a" "n < 0" shows "power_int a n \ power_int b n" using power_int_strict_antimono[of a b n] assms by (cases "a = b") auto end subsection \Finiteness of intervals\ lemma finite_interval_int1 [iff]: "finite {i :: int. a \ i \ i \ b}" proof (cases "a \ b") case True then show ?thesis proof (induct b rule: int_ge_induct) case base have "{i. a \ i \ i \ a} = {a}" by auto then show ?case by simp next case (step b) then have "{i. a \ i \ i \ b + 1} = {i. a \ i \ i \ b} \ {b + 1}" by auto with step show ?case by simp qed next case False then show ?thesis by (metis (lifting, no_types) Collect_empty_eq finite.emptyI order_trans) qed lemma finite_interval_int2 [iff]: "finite {i :: int. a \ i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \ i \ b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto subsection \Configuration of the code generator\ text \Constructors\ definition Pos :: "num \ int" where [simp, code_abbrev]: "Pos = numeral" definition Neg :: "num \ int" where [simp, code_abbrev]: "Neg n = - (Pos n)" code_datatype "0::int" Pos Neg text \Auxiliary operations.\ definition dup :: "int \ int" where [simp]: "dup k = k + k" lemma dup_code [code]: "dup 0 = 0" "dup (Pos n) = Pos (Num.Bit0 n)" "dup (Neg n) = Neg (Num.Bit0 n)" by (simp_all add: numeral_Bit0) definition sub :: "num \ num \ int" where [simp]: "sub m n = numeral m - numeral n" lemma sub_code [code]: "sub Num.One Num.One = 0" "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM) lemma sub_BitM_One_eq: \(Num.sub (Num.BitM n) num.One) = 2 * (Num.sub n Num.One :: int)\ by (cases n) simp_all text \Implementations.\ lemma one_int_code [code]: "1 = Pos Num.One" by simp lemma plus_int_code [code]: "k + 0 = k" "0 + l = l" "Pos m + Pos n = Pos (m + n)" "Pos m + Neg n = sub m n" "Neg m + Pos n = sub n m" "Neg m + Neg n = Neg (m + n)" for k l :: int by simp_all lemma uminus_int_code [code]: "uminus 0 = (0::int)" "uminus (Pos m) = Neg m" "uminus (Neg m) = Pos m" by simp_all lemma minus_int_code [code]: "k - 0 = k" "0 - l = uminus l" "Pos m - Pos n = sub m n" "Pos m - Neg n = Pos (m + n)" "Neg m - Pos n = Neg (m + n)" "Neg m - Neg n = sub n m" for k l :: int by simp_all lemma times_int_code [code]: "k * 0 = 0" "0 * l = 0" "Pos m * Pos n = Pos (m * n)" "Pos m * Neg n = Neg (m * n)" "Neg m * Pos n = Neg (m * n)" "Neg m * Neg n = Pos (m * n)" for k l :: int by simp_all instantiation int :: equal begin definition "HOL.equal k l \ k = (l::int)" instance by standard (rule equal_int_def) end lemma equal_int_code [code]: "HOL.equal 0 (0::int) \ True" "HOL.equal 0 (Pos l) \ False" "HOL.equal 0 (Neg l) \ False" "HOL.equal (Pos k) 0 \ False" "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" "HOL.equal (Pos k) (Neg l) \ False" "HOL.equal (Neg k) 0 \ False" "HOL.equal (Neg k) (Pos l) \ False" "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" by (auto simp add: equal) lemma equal_int_refl [code nbe]: "HOL.equal k k \ True" for k :: int by (fact equal_refl) lemma less_eq_int_code [code]: "0 \ (0::int) \ True" "0 \ Pos l \ True" "0 \ Neg l \ False" "Pos k \ 0 \ False" "Pos k \ Pos l \ k \ l" "Pos k \ Neg l \ False" "Neg k \ 0 \ True" "Neg k \ Pos l \ True" "Neg k \ Neg l \ l \ k" by simp_all lemma less_int_code [code]: "0 < (0::int) \ False" "0 < Pos l \ True" "0 < Neg l \ False" "Pos k < 0 \ False" "Pos k < Pos l \ k < l" "Pos k < Neg l \ False" "Neg k < 0 \ True" "Neg k < Pos l \ True" "Neg k < Neg l \ l < k" by simp_all lemma nat_code [code]: "nat (Int.Neg k) = 0" "nat 0 = 0" "nat (Int.Pos k) = nat_of_num k" by (simp_all add: nat_of_num_numeral) lemma (in ring_1) of_int_code [code]: "of_int (Int.Neg k) = - numeral k" "of_int 0 = 0" "of_int (Int.Pos k) = numeral k" by simp_all text \Serializer setup.\ code_identifier code_module Int \ (SML) Arith and (OCaml) Arith and (Haskell) Arith quickcheck_params [default_type = int] hide_const (open) Pos Neg sub dup text \De-register \int\ as a quotient type:\ lifting_update int.lifting lifting_forget int.lifting subsection \Duplicates\ lemmas int_sum = of_nat_sum [where 'a=int] lemmas int_prod = of_nat_prod [where 'a=int] lemmas zle_int = of_nat_le_iff [where 'a=int] lemmas int_int_eq = of_nat_eq_iff [where 'a=int] lemmas nonneg_eq_int = nonneg_int_cases lemmas double_eq_0_iff = double_zero lemmas int_distrib = distrib_right [of z1 z2 w] distrib_left [of w z1 z2] left_diff_distrib [of z1 z2 w] right_diff_distrib [of w z1 z2] for z1 z2 w :: int 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,302 +1,299 @@ (* Author: Various *) section \Combination and Cancellation Simprocs for Numeral Expressions\ theory Numeral_Simprocs imports Divides 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 -declare split_div [of _ _ "numeral k", linarith_split] for k -declare split_mod [of _ _ "numeral k", linarith_split] for k - 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/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2759 +1,2785 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all lemma split_of_bool_asm: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all lemma of_bool_eq_0_iff [simp]: \of_bool P = 0 \ \ P\ by simp lemma of_bool_eq_1_iff [simp]: \of_bool P = 1 \ P\ by simp end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" using that by (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. lemma of_bool_not_iff [simp]: \of_bool (\ P) = 1 - of_bool P\ by simp lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed +lemma inj_mult_left [simp]: \inj ((*) a) \ a \ 0\ (is \?P \ ?Q\) +proof + assume ?P + show ?Q + proof + assume \a = 0\ + with \?P\ have "inj ((*) 0)" + by simp + moreover have "0 * 0 = 0 * 1" + by simp + ultimately have "0 = 1" + by (rule injD) + then show False + by simp + qed +next + assume ?Q then show ?P + by (auto intro: injI) +qed + end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" using assms by (simp add: dvd_div_eq_0_iff unit_imp_dvd) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ begin lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right end class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp lemma normalize_mult_unit_left [simp]: assumes "a dvd 1" shows "normalize (a * b) = normalize b" proof (cases "b = 0") case False have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" using assms by (subst unit_factor_mult_unit_left) auto also have "\ = a * b" by simp also have "b = unit_factor b * normalize b" by simp hence "a * b = a * unit_factor b * normalize b" by (simp only: mult_ac) finally show ?thesis using assms False by auto qed auto lemma normalize_mult_unit_right [simp]: assumes "b dvd 1" shows "normalize (a * b) = normalize a" using assms by (subst mult.commute) auto lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") case False have "normalize a = normalize (unit_factor a * normalize a)" by simp also from False have "\ = normalize (normalize a)" by (subst normalize_mult_unit_left) auto finally show ?thesis .. qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" by (rule iffI; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) end class normalization_semidom_multiplicative = normalization_semidom + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" begin lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False have "unit_factor (a * b) * normalize (a * b) = a * b" by (rule unit_factor_mult_normalize) then have "normalize (a * b) = a * b div unit_factor (a * b)" by simp also have "\ = a * b div unit_factor (b * a)" by (simp add: ac_simps) also have "\ = a * b div unit_factor b div unit_factor a" using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) also have "\ = a * (b div unit_factor b) div unit_factor a" using False by (subst unit_div_mult_swap) simp_all also have "\ = normalize a * normalize b" using False by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) finally show ?thesis . qed lemma dvd_unit_factor_div: assumes "b dvd a" shows "unit_factor (a div b) = unit_factor a div unit_factor b" proof - from assms have "a = a div b * b" by simp then have "unit_factor a = unit_factor (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: unit_factor_mult) qed lemma dvd_normalize_div: assumes "b dvd a" shows "normalize (a div b) = normalize a div normalize b" proof - from assms have "a = a div b * b" by simp then have "normalize a = normalize (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: normalize_mult) qed end text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end lemma mono_mult: fixes a :: "'a::ordered_semiring" shows "a \ 0 \ mono ((*) a)" by (simp add: mono_def mult_left_mono) class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" begin subclass zero_neq_one by standard (simp add: less_imp_neq) lemma zero_le_one [simp]: \0 \ 1\ by (rule less_imp_le) simp end class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: assumes "0 < a * b" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b]) qed (auto simp add: le_less not_less) lemma zero_less_mult_pos2: assumes "0 < b * a" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b]) qed (auto simp add: le_less not_less) text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" "c < d" "0 < b" "0 \ c" shows "a * c < b * d" proof (cases "c = 0") case True with assms show ?thesis by simp next case False with assms have "a*c < b*c" by (simp add: mult_strict_right_mono [OF \a < b\]) also have "\ < b*d" by (simp add: assms mult_strict_left_mono) finally show ?thesis . qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" using assms by (auto simp add: mult_strict_mono) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" proof - have "a * c < b * c" by (simp add: assms mult_strict_right_mono) also have "... \ b * d" by (intro mult_left_mono) (use assms in auto) finally show ?thesis . qed lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" proof - have "a * c \ b * c" by (simp add: assms mult_right_mono) also have "... < b * d" by (intro mult_strict_left_mono) (use assms in auto) finally show ?thesis . qed end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" by (auto dest: mult_left_mono [of _ _ "- c"]) lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" by (auto dest: mult_right_mono [of _ _ "- c"]) lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg) qed qed auto lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg) qed qed auto text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one by standard subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) lemma of_bool_less_eq_iff [simp]: \of_bool P \ of_bool Q \ (P \ Q)\ by auto lemma of_bool_less_iff [simp]: \of_bool P < of_bool Q \ \ P \ Q\ by auto lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b proof (rule ccontr) assume "\ a + 1 < b + 1" moreover with that have "a + 1 < b + 1" by simp ultimately show False by contradiction qed qed lemma zero_less_eq_of_bool [simp]: \0 \ of_bool P\ by simp lemma zero_less_of_bool_iff [simp]: \0 < of_bool P \ P\ by simp lemma of_bool_less_eq_one [simp]: \of_bool P \ 1\ by simp lemma of_bool_less_one_iff [simp]: \of_bool P < 1 \ \ P\ by simp lemma of_bool_or_iff [simp]: \of_bool (P \ Q) = max (of_bool P) (of_bool Q)\ by (simp add: max_def) text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: assumes "i + k \ n" shows "i \ n - k" proof - have "n - (i + k) + i + k = n" by (simp add: assms add.assoc) with assms add_implies_diff have "i + k \ n - k + k" by fastforce then show ?thesis by simp qed lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + i + k = n" using 1 by (simp add: add.assoc) moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 add_le_imp_le_diff [of "n-k" k "j + k"] by (simp add: add.commute diff_diff_add) qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) lemma abs_bool_eq [simp]: \\of_bool P\ = of_bool P\ by simp lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma left_sgn_mult_self_eq [simp]: \sgn a * (sgn a * b) = of_bool (a \ 0) * b\ by (simp flip: mult.assoc) lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) +text \ + Is this really better than just rewriting with \abs_if\? +\ +lemma abs_split [no_atp]: \P \a\ \ (0 \ a \ P a) \ (a < 0 \ P (- a))\ + by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) + end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_mult_pos': "0 \ x \ x * \y\ = \x * y\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end