# HG changeset patch # User haftmann # Date 1265645558 -3600 # Node ID 9f841f20dca6798de52437e34dff846beda2f7cf # Parent 00f311c324446810f13723c6c5372f90be721038 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields diff -r 00f311c32444 -r 9f841f20dca6 NEWS --- a/NEWS Mon Feb 08 17:12:32 2010 +0100 +++ b/NEWS Mon Feb 08 17:12:38 2010 +0100 @@ -15,6 +15,11 @@ * New set of rules "ac_simps" provides combined assoc / commute rewrites for all interpretations of the appropriate generic locales. +* Renamed theory "OrderedGroup" to "Groups" and split theory "Ring_and_Field" +into theories "Rings" and "Fields"; for more appropriate and more +consistent names suitable for name prefixes within the HOL theories. +INCOMPATIBILITY. + * More consistent naming of type classes involving orderings (and lattices): lower_semilattice ~> semilattice_inf diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Algebra/poly/UnivPoly2.thy --- a/src/HOL/Algebra/poly/UnivPoly2.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Algebra/poly/UnivPoly2.thy Mon Feb 08 17:12:38 2010 +0100 @@ -139,7 +139,7 @@ end -instance up :: ("{times, comm_monoid_add}") Ring_and_Field.dvd .. +instance up :: ("{times, comm_monoid_add}") Rings.dvd .. instantiation up :: ("{times, one, comm_monoid_add, uminus, minus}") inverse begin diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Decision_Procs/mir_tac.ML Mon Feb 08 17:12:38 2010 +0100 @@ -33,7 +33,7 @@ @{thm "real_of_nat_number_of"}, @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"}, @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"}, - @{thm "Ring_and_Field.divide_zero"}, + @{thm "Fields.divide_zero"}, @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"}, @{thm "diff_def"}, @{thm "minus_divide_left"}] diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Divides.thy Mon Feb 08 17:12:38 2010 +0100 @@ -657,7 +657,7 @@ val trans = trans; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac - (@{thm monoid_add_class.add_0_left} :: @{thm monoid_add_class.add_0_right} :: @{thms add_ac})) + (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac})) end) @@ -1655,8 +1655,8 @@ lemmas arithmetic_simps = arith_simps add_special - OrderedGroup.add_0_left - OrderedGroup.add_0_right + add_0_left + add_0_right mult_zero_left mult_zero_right mult_1_left diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Fields.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Fields.thy Mon Feb 08 17:12:38 2010 +0100 @@ -0,0 +1,1044 @@ +(* Title: HOL/Fields.thy + Author: Gertrud Bauer + Author: Steven Obua + Author: Tobias Nipkow + Author: Lawrence C Paulson + Author: Markus Wenzel + Author: Jeremy Avigad +*) + +header {* Fields *} + +theory Fields +imports Rings +begin + +class field = comm_ring_1 + inverse + + assumes field_inverse: "a \ 0 \ inverse a * a = 1" + assumes divide_inverse: "a / b = a * inverse b" +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) +qed + +subclass idom .. + +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_ac) + 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 divide_zero_left [simp]: "0 / a = 0" +by (simp add: divide_inverse) + +lemma inverse_eq_divide: "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) + +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 mult_ac) + +lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]: +assumes [simp]: "b\0" and [simp]: "c\0" shows "(c*a)/(c*b) = a/b" +proof - + 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: mult_ac) + also have "... = a * inverse b" by simp + finally show ?thesis by (simp add: divide_inverse) +qed + +lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]: + "\b \ 0; c \ 0\ \ (a * c) / (b * c) = a / b" +by (simp add: mult_commute [of _ c]) + +lemma divide_1 [simp]: "a / 1 = a" +by (simp add: divide_inverse) + +lemma times_divide_eq_right: "a * (b / c) = (a * b) / c" +by (simp add: divide_inverse mult_assoc) + +lemma times_divide_eq_left: "(b / c) * a = (b * a) / c" +by (simp add: divide_inverse mult_ac) + +text {* These are later declared as simp rules. *} +lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left + +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_mult_divide_cancel_right [simp, noatp]: + "b \ 0 \ a * b / b = a" +using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp + +lemma nonzero_mult_divide_cancel_left [simp, noatp]: + "a \ 0 \ a * b / a = b" +using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp + +lemma nonzero_divide_mult_cancel_right [simp, noatp]: + "\a \ 0; b \ 0\ \ b / (a * b) = 1 / a" +using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp + +lemma nonzero_divide_mult_cancel_left [simp, noatp]: + "\a \ 0; b \ 0\ \ a / (a * b) = 1 / b" +using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp + +lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]: + "\b \ 0; c \ 0\ \ (c * a) / (b * c) = a / b" +using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac) + +lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]: + "\b \ 0; c \ 0\ \ (a * c) / (c * b) = a / b" +using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac) + +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, noatp]: "(-a) / b = - (a / b)" +by (simp add: divide_inverse) + +lemma diff_divide_distrib: "(a - b) / c = a / c - b / c" +by (simp add: diff_minus add_divide_distrib) + +lemma add_divide_eq_iff: + "z \ 0 \ x + y / z = (z * x + y) / z" +by (simp add: add_divide_distrib) + +lemma divide_add_eq_iff: + "z \ 0 \ x / z + y = (x + z * y) / z" +by (simp add: add_divide_distrib) + +lemma diff_divide_eq_iff: + "z \ 0 \ x - y / z = (z * x - y) / z" +by (simp add: diff_divide_distrib) + +lemma divide_diff_eq_iff: + "z \ 0 \ x / z - y = (x - z * y) / z" +by (simp add: diff_divide_distrib) + +lemma nonzero_eq_divide_eq: "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: "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 divide_eq_imp: "c \ 0 \ b = a * c \ b / c = a" +by simp + +lemma eq_divide_imp: "c \ 0 \ a * c = b \ a = b / c" +by (erule subst, simp) + +lemmas field_eq_simps[noatp] = algebra_simps + (* pull / out*) + add_divide_eq_iff divide_add_eq_iff + diff_divide_eq_iff divide_diff_eq_iff + (* multiply eqn *) + nonzero_eq_divide_eq nonzero_divide_eq_eq +(* is added later: + times_divide_eq_left times_divide_eq_right +*) + +text{*An example:*} +lemma "\a\b; c\d; e\f\ \ ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1" +apply(subgoal_tac "(c-d)*(e-f)*(a-b) \ 0") + apply(simp add:field_eq_simps) +apply(simp) +done + +lemma diff_frac_eq: + "y \ 0 \ z \ 0 \ x / y - w / z = (x * z - w * y) / (y * z)" +by (simp add: field_eq_simps times_divide_eq) + +lemma frac_eq_eq: + "y \ 0 \ z \ 0 \ (x / y = w / z) = (x * z = w * y)" +by (simp add: field_eq_simps times_divide_eq) + +end + +class division_by_zero = zero + inverse + + assumes inverse_zero [simp]: "inverse 0 = 0" + +lemma divide_zero [simp]: + "a / 0 = (0::'a::{field,division_by_zero})" +by (simp add: divide_inverse) + +lemma divide_self_if [simp]: + "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)" +by simp + +class linordered_field = field + linordered_idom + +lemma inverse_nonzero_iff_nonzero [simp]: + "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))" +by (force dest: inverse_zero_imp_zero) + +lemma inverse_minus_eq [simp]: + "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})" +proof cases + assume "a=0" thus ?thesis by (simp add: inverse_zero) +next + assume "a\0" + thus ?thesis by (simp add: nonzero_inverse_minus_eq) +qed + +lemma inverse_eq_imp_eq: + "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})" +apply (cases "a=0 | b=0") + apply (force dest!: inverse_zero_imp_zero + simp add: eq_commute [of "0::'a"]) +apply (force dest!: nonzero_inverse_eq_imp_eq) +done + +lemma inverse_eq_iff_eq [simp]: + "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))" +by (force dest!: inverse_eq_imp_eq) + +lemma inverse_inverse_eq [simp]: + "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a" + proof cases + assume "a=0" thus ?thesis by simp + next + assume "a\0" + thus ?thesis by (simp add: nonzero_inverse_inverse_eq) + qed + +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::'a::{field,division_by_zero})" + proof cases + assume "a \ 0 & b \ 0" + thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute) + next + assume "~ (a \ 0 & b \ 0)" + thus ?thesis by force + qed + +lemma inverse_divide [simp]: + "inverse (a/b) = b / (a::'a::{field,division_by_zero})" +by (simp add: divide_inverse mult_commute) + + +subsection {* Calculations with fractions *} + +text{* There is a whole bunch of simp-rules just for class @{text +field} but none for class @{text field} and @{text nonzero_divides} +because the latter are covered by a simproc. *} + +lemma mult_divide_mult_cancel_left: + "c\0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})" +apply (cases "b = 0") +apply (simp_all add: nonzero_mult_divide_mult_cancel_left) +done + +lemma mult_divide_mult_cancel_right: + "c\0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})" +apply (cases "b = 0") +apply (simp_all add: nonzero_mult_divide_mult_cancel_right) +done + +lemma divide_divide_eq_right [simp,noatp]: + "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})" +by (simp add: divide_inverse mult_ac) + +lemma divide_divide_eq_left [simp,noatp]: + "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)" +by (simp add: divide_inverse mult_assoc) + + +subsubsection{*Special Cancellation Simprules for Division*} + +lemma mult_divide_mult_cancel_left_if[simp,noatp]: +fixes c :: "'a :: {field,division_by_zero}" +shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)" +by (simp add: mult_divide_mult_cancel_left) + + +subsection {* Division and Unary Minus *} + +lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})" +by (simp add: divide_inverse) + +lemma divide_minus_right [simp, noatp]: + "a / -(b::'a::{field,division_by_zero}) = -(a / b)" +by (simp add: divide_inverse) + +lemma minus_divide_divide: + "(-a)/(-b) = a / (b::'a::{field,division_by_zero})" +apply (cases "b=0", simp) +apply (simp add: nonzero_minus_divide_divide) +done + +lemma eq_divide_eq: + "((a::'a::{field,division_by_zero}) = b/c) = (if c\0 then a*c = b else a=0)" +by (simp add: nonzero_eq_divide_eq) + +lemma divide_eq_eq: + "(b/c = (a::'a::{field,division_by_zero})) = (if c\0 then b = a*c else a=0)" +by (force simp add: nonzero_divide_eq_eq) + + +subsection {* Ordered Fields *} + +lemma positive_imp_inverse_positive: +assumes a_gt_0: "0 < a" shows "0 < inverse (a::'a::linordered_field)" +proof - + have "0 < a * inverse a" + by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one) + thus "0 < inverse a" + by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff) +qed + +lemma negative_imp_inverse_negative: + "a < 0 ==> inverse a < (0::'a::linordered_field)" +by (insert positive_imp_inverse_positive [of "-a"], + simp add: nonzero_inverse_minus_eq order_less_imp_not_eq) + +lemma inverse_le_imp_le: +assumes invle: "inverse a \ inverse b" and apos: "0 < a" +shows "b \ (a::'a::linordered_field)" +proof (rule classical) + assume "~ b \ a" + hence "a < b" by (simp add: linorder_not_le) + hence bpos: "0 < b" by (blast intro: apos order_less_trans) + hence "a * inverse a \ a * inverse b" + by (simp add: apos invle order_less_imp_le mult_left_mono) + hence "(a * inverse a) * b \ (a * inverse b) * b" + by (simp add: bpos order_less_imp_le mult_right_mono) + thus "b \ a" by (simp add: mult_assoc apos bpos order_less_imp_not_eq2) +qed + +lemma inverse_positive_imp_positive: +assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" +shows "0 < (a::'a::linordered_field)" +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_positive_iff_positive [simp]: + "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))" +apply (cases "a = 0", simp) +apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive) +done + +lemma inverse_negative_imp_negative: +assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" +shows "a < (0::'a::linordered_field)" +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 inverse_negative_iff_negative [simp]: + "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))" +apply (cases "a = 0", simp) +apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative) +done + +lemma inverse_nonnegative_iff_nonnegative [simp]: + "(0 \ inverse a) = (0 \ (a::'a::{linordered_field,division_by_zero}))" +by (simp add: linorder_not_less [symmetric]) + +lemma inverse_nonpositive_iff_nonpositive [simp]: + "(inverse a \ 0) = (a \ (0::'a::{linordered_field,division_by_zero}))" +by (simp add: linorder_not_less [symmetric]) + +lemma linordered_field_no_lb: "\ x. \y. y < (x::'a::linordered_field)" +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::'a::linordered_field)" +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 + +subsection{*Anti-Monotonicity of @{term inverse}*} + +lemma less_imp_inverse_less: +assumes less: "a < b" and apos: "0 < a" +shows "inverse b < inverse (a::'a::linordered_field)" +proof (rule ccontr) + assume "~ inverse b < inverse a" + hence "inverse a \ inverse b" by (simp add: linorder_not_less) + hence "~ (a < b)" + by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos]) + thus False by (rule notE [OF _ less]) +qed + +lemma inverse_less_imp_less: + "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)" +apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"]) +apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) +done + +text{*Both premises are essential. Consider -1 and 1.*} +lemma inverse_less_iff_less [simp,noatp]: + "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))" +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::'a::linordered_field)" +by (force simp add: order_le_less less_imp_inverse_less) + +lemma inverse_le_iff_le [simp,noatp]: + "[|0 < a; 0 < b|] ==> (inverse a \ inverse b) = (b \ (a::'a::linordered_field))" +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: + "[|inverse a \ inverse b; b < 0|] ==> b \ (a::'a::linordered_field)" +apply (rule classical) +apply (subgoal_tac "a < 0") + prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) +apply (insert inverse_le_imp_le [of "-b" "-a"]) +apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) +done + +lemma less_imp_inverse_less_neg: + "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)" +apply (subgoal_tac "a < 0") + prefer 2 apply (blast intro: order_less_trans) +apply (insert less_imp_inverse_less [of "-b" "-a"]) +apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) +done + +lemma inverse_less_imp_less_neg: + "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)" +apply (rule classical) +apply (subgoal_tac "a < 0") + prefer 2 + apply (force simp add: linorder_not_less intro: order_le_less_trans) +apply (insert inverse_less_imp_less [of "-b" "-a"]) +apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) +done + +lemma inverse_less_iff_less_neg [simp,noatp]: + "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))" +apply (insert inverse_less_iff_less [of "-b" "-a"]) +apply (simp del: inverse_less_iff_less + add: order_less_imp_not_eq nonzero_inverse_minus_eq) +done + +lemma le_imp_inverse_le_neg: + "[|a \ b; b < 0|] ==> inverse b \ inverse (a::'a::linordered_field)" +by (force simp add: order_le_less less_imp_inverse_less_neg) + +lemma inverse_le_iff_le_neg [simp,noatp]: + "[|a < 0; b < 0|] ==> (inverse a \ inverse b) = (b \ (a::'a::linordered_field))" +by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) + + +subsection{*Inverses and the Number One*} + +lemma one_less_inverse_iff: + "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))" +proof cases + assume "0 < x" + with inverse_less_iff_less [OF zero_less_one, of x] + show ?thesis by simp +next + assume notless: "~ (0 < x)" + have "~ (1 < inverse x)" + proof + assume "1 < inverse x" + also with notless have "... \ 0" by (simp add: linorder_not_less) + also have "... < 1" by (rule zero_less_one) + finally show False by auto + qed + with notless show ?thesis by simp +qed + +lemma inverse_eq_1_iff [simp]: + "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))" +by (insert inverse_eq_iff_eq [of x 1], simp) + +lemma one_le_inverse_iff: + "(1 \ inverse x) = (0 < x & x \ (1::'a::{linordered_field,division_by_zero}))" +by (force simp add: order_le_less one_less_inverse_iff zero_less_one + eq_commute [of 1]) + +lemma inverse_less_1_iff: + "(inverse x < 1) = (x \ 0 | 1 < (x::'a::{linordered_field,division_by_zero}))" +by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) + +lemma inverse_le_1_iff: + "(inverse x \ 1) = (x \ 0 | 1 \ (x::'a::{linordered_field,division_by_zero}))" +by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) + + +subsection{*Simplification of Inequalities Involving Literal Divisors*} + +lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \ b/c) = (a*c \ b)" +proof - + assume less: "0 b/c) = (a*c \ (b/c)*c)" + by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) + also have "... = (a*c \ b)" + by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \ b/c) = (b \ a*c)" +proof - + assume less: "c<0" + hence "(a \ b/c) = ((b/c)*c \ a*c)" + by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) + also have "... = (b \ a*c)" + by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma le_divide_eq: + "(a \ b/c) = + (if 0 < c then a*c \ b + else if c < 0 then b \ a*c + else a \ (0::'a::{linordered_field,division_by_zero}))" +apply (cases "c=0", simp) +apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) +done + +lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \ a) = (b \ a*c)" +proof - + assume less: "0 a) = ((b/c)*c \ a*c)" + by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) + also have "... = (b \ a*c)" + by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \ a) = (a*c \ b)" +proof - + assume less: "c<0" + hence "(b/c \ a) = (a*c \ (b/c)*c)" + by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) + also have "... = (a*c \ b)" + by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma divide_le_eq: + "(b/c \ a) = + (if 0 < c then b \ a*c + else if c < 0 then a*c \ b + else 0 \ (a::'a::{linordered_field,division_by_zero}))" +apply (cases "c=0", simp) +apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) +done + +lemma pos_less_divide_eq: + "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)" +proof - + assume less: "0 (a < b/c) = (b < a*c)" +proof - + assume less: "c<0" + hence "(a < b/c) = ((b/c)*c < a*c)" + by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) + also have "... = (b < a*c)" + by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma less_divide_eq: + "(a < b/c) = + (if 0 < c then a*c < b + else if c < 0 then b < a*c + else a < (0::'a::{linordered_field,division_by_zero}))" +apply (cases "c=0", simp) +apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) +done + +lemma pos_divide_less_eq: + "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)" +proof - + assume less: "0 (b/c < a) = (a*c < b)" +proof - + assume less: "c<0" + hence "(b/c < a) = (a*c < (b/c)*c)" + by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) + also have "... = (a*c < b)" + by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) + finally show ?thesis . +qed + +lemma divide_less_eq: + "(b/c < a) = + (if 0 < c then b < a*c + else if c < 0 then a*c < b + else 0 < (a::'a::{linordered_field,division_by_zero}))" +apply (cases "c=0", simp) +apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) +done + + +subsection{*Field simplification*} + +text{* Lemmas @{text field_simps} multiply with denominators in in(equations) +if they can be proved to be non-zero (for equations) or positive/negative +(for inequations). Can be too aggressive and is therefore separate from the +more benign @{text algebra_simps}. *} + +lemmas field_simps[noatp] = field_eq_simps + (* multiply ineqn *) + pos_divide_less_eq neg_divide_less_eq + pos_less_divide_eq neg_less_divide_eq + pos_divide_le_eq neg_divide_le_eq + pos_le_divide_eq neg_le_divide_eq + +text{* Lemmas @{text sign_simps} is a first attempt to automate proofs +of positivity/negativity needed for @{text field_simps}. Have not added @{text +sign_simps} to @{text field_simps} because the former can lead to case +explosions. *} + +lemmas sign_simps[noatp] = group_simps + zero_less_mult_iff mult_less_0_iff + +(* Only works once linear arithmetic is installed: +text{*An example:*} +lemma fixes a b c d e f :: "'a::linordered_field" +shows "\a>b; c \ + ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) < + ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u" +apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0") + prefer 2 apply(simp add:sign_simps) +apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0") + prefer 2 apply(simp add:sign_simps) +apply(simp add:field_simps) +done +*) + + +subsection{*Division and Signs*} + +lemma zero_less_divide_iff: + "((0::'a::{linordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)" +by (simp add: divide_inverse zero_less_mult_iff) + +lemma divide_less_0_iff: + "(a/b < (0::'a::{linordered_field,division_by_zero})) = + (0 < a & b < 0 | a < 0 & 0 < b)" +by (simp add: divide_inverse mult_less_0_iff) + +lemma zero_le_divide_iff: + "((0::'a::{linordered_field,division_by_zero}) \ a/b) = + (0 \ a & 0 \ b | a \ 0 & b \ 0)" +by (simp add: divide_inverse zero_le_mult_iff) + +lemma divide_le_0_iff: + "(a/b \ (0::'a::{linordered_field,division_by_zero})) = + (0 \ a & b \ 0 | a \ 0 & 0 \ b)" +by (simp add: divide_inverse mult_le_0_iff) + +lemma divide_eq_0_iff [simp,noatp]: + "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))" +by (simp add: divide_inverse) + +lemma divide_pos_pos: + "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y" +by(simp add:field_simps) + + +lemma divide_nonneg_pos: + "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y" +by(simp add:field_simps) + +lemma divide_neg_pos: + "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0" +by(simp add:field_simps) + +lemma divide_nonpos_pos: + "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0" +by(simp add:field_simps) + +lemma divide_pos_neg: + "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0" +by(simp add:field_simps) + +lemma divide_nonneg_neg: + "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0" +by(simp add:field_simps) + +lemma divide_neg_neg: + "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y" +by(simp add:field_simps) + +lemma divide_nonpos_neg: + "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y" +by(simp add:field_simps) + + +subsection{*Cancellation Laws for Division*} + +lemma divide_cancel_right [simp,noatp]: + "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))" +apply (cases "c=0", simp) +apply (simp add: divide_inverse) +done + +lemma divide_cancel_left [simp,noatp]: + "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" +apply (cases "c=0", simp) +apply (simp add: divide_inverse) +done + + +subsection {* Division and the Number One *} + +text{*Simplify expressions equated with 1*} +lemma divide_eq_1_iff [simp,noatp]: + "(a/b = 1) = (b \ 0 & a = (b::'a::{field,division_by_zero}))" +apply (cases "b=0", simp) +apply (simp add: right_inverse_eq) +done + +lemma one_eq_divide_iff [simp,noatp]: + "(1 = a/b) = (b \ 0 & a = (b::'a::{field,division_by_zero}))" +by (simp add: eq_commute [of 1]) + +lemma zero_eq_1_divide_iff [simp,noatp]: + "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)" +apply (cases "a=0", simp) +apply (auto simp add: nonzero_eq_divide_eq) +done + +lemma one_divide_eq_0_iff [simp,noatp]: + "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)" +apply (cases "a=0", simp) +apply (insert zero_neq_one [THEN not_sym]) +apply (auto simp add: nonzero_divide_eq_eq) +done + +text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*} +lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified] +lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified] +lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified] +lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified] + +declare zero_less_divide_1_iff [simp,noatp] +declare divide_less_0_1_iff [simp,noatp] +declare zero_le_divide_1_iff [simp,noatp] +declare divide_le_0_1_iff [simp,noatp] + + +subsection {* Ordering Rules for Division *} + +lemma divide_strict_right_mono: + "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)" +by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono + positive_imp_inverse_positive) + +lemma divide_right_mono: + "[|a \ b; 0 \ c|] ==> a/c \ b/(c::'a::{linordered_field,division_by_zero})" +by (force simp add: divide_strict_right_mono order_le_less) + +lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b + ==> c <= 0 ==> b / c <= a / c" +apply (drule divide_right_mono [of _ _ "- c"]) +apply auto +done + +lemma divide_strict_right_mono_neg: + "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)" +apply (drule divide_strict_right_mono [of _ _ "-c"], simp) +apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric]) +done + +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::'a::linordered_field)" +by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono) + +lemma divide_left_mono: + "[|b \ a; 0 \ c; 0 < a*b|] ==> c / a \ c / (b::'a::linordered_field)" +by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono) + +lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b + ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b" + apply (drule divide_left_mono [of _ _ "- c"]) + apply (auto simp add: mult_commute) +done + +lemma divide_strict_left_mono_neg: + "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)" +by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg) + + +text{*Simplify quotients that are compared with the value 1.*} + +lemma le_divide_eq_1 [noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(1 \ b / a) = ((0 < a & a \ b) | (a < 0 & b \ a))" +by (auto simp add: le_divide_eq) + +lemma divide_le_eq_1 [noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(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 [noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))" +by (auto simp add: less_divide_eq) + +lemma divide_less_eq_1 [noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)" +by (auto simp add: divide_less_eq) + + +subsection{*Conditional Simplification Rules: No Case Splits*} + +lemma le_divide_eq_1_pos [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "0 < a \ (1 \ b/a) = (a \ b)" +by (auto simp add: le_divide_eq) + +lemma le_divide_eq_1_neg [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "a < 0 \ (1 \ b/a) = (b \ a)" +by (auto simp add: le_divide_eq) + +lemma divide_le_eq_1_pos [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "0 < a \ (b/a \ 1) = (b \ a)" +by (auto simp add: divide_le_eq) + +lemma divide_le_eq_1_neg [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "a < 0 \ (b/a \ 1) = (a \ b)" +by (auto simp add: divide_le_eq) + +lemma less_divide_eq_1_pos [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "0 < a \ (1 < b/a) = (a < b)" +by (auto simp add: less_divide_eq) + +lemma less_divide_eq_1_neg [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "a < 0 \ (1 < b/a) = (b < a)" +by (auto simp add: less_divide_eq) + +lemma divide_less_eq_1_pos [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "0 < a \ (b/a < 1) = (b < a)" +by (auto simp add: divide_less_eq) + +lemma divide_less_eq_1_neg [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "a < 0 \ b/a < 1 <-> a < b" +by (auto simp add: divide_less_eq) + +lemma eq_divide_eq_1 [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(1 = b/a) = ((a \ 0 & a = b))" +by (auto simp add: eq_divide_eq) + +lemma divide_eq_eq_1 [simp,noatp]: + fixes a :: "'a :: {linordered_field,division_by_zero}" + shows "(b/a = 1) = ((a \ 0 & a = b))" +by (auto simp add: divide_eq_eq) + + +subsection {* Reasoning about inequalities with division *} + +lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==> + x / y <= z" +by (subst pos_divide_le_eq, assumption+) + +lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==> + z <= x / y" +by(simp add:field_simps) + +lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==> + x / y < z" +by(simp add:field_simps) + +lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==> + z < x / y" +by(simp add:field_simps) + +lemma frac_le: "(0::'a::linordered_field) <= x ==> + x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w" + apply (rule mult_imp_div_pos_le) + apply simp + apply (subst times_divide_eq_left) + apply (rule mult_imp_le_div_pos, assumption) + apply (rule mult_mono) + apply simp_all +done + +lemma frac_less: "(0::'a::linordered_field) <= x ==> + x < y ==> 0 < w ==> w <= z ==> x / z < y / w" + apply (rule mult_imp_div_pos_less) + apply simp + apply (subst times_divide_eq_left) + apply (rule mult_imp_less_div_pos, assumption) + apply (erule mult_less_le_imp_less) + apply simp_all +done + +lemma frac_less2: "(0::'a::linordered_field) < x ==> + x <= y ==> 0 < w ==> w < z ==> x / z < y / w" + apply (rule mult_imp_div_pos_less) + apply simp_all + apply (subst times_divide_eq_left) + apply (rule mult_imp_less_div_pos, assumption) + apply (erule mult_le_less_imp_less) + apply simp_all +done + +text{*It's not obvious whether these should be simprules or not. + Their effect is to gather terms into one big fraction, like + a*b*c / x*y*z. The rationale for that is unclear, but many proofs + seem to need them.*} + +declare times_divide_eq [simp] + + +subsection {* Ordered Fields are Dense *} + +lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)" +by (simp add: field_simps zero_less_two) + +lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b" +by (simp add: field_simps zero_less_two) + +instance linordered_field < dense_linorder +proof + fix x y :: 'a + have "x < x + 1" by simp + then show "\y. x < y" .. + have "x - 1 < x" by simp + then show "\y. y < x" .. + show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) +qed + + +subsection {* Absolute Value *} + +lemma nonzero_abs_inverse: + "a \ 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)" +apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq + negative_imp_inverse_negative) +apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) +done + +lemma abs_inverse [simp]: + "abs (inverse (a::'a::{linordered_field,division_by_zero})) = + inverse (abs a)" +apply (cases "a=0", simp) +apply (simp add: nonzero_abs_inverse) +done + +lemma nonzero_abs_divide: + "b \ 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b" +by (simp add: divide_inverse abs_mult nonzero_abs_inverse) + +lemma abs_divide [simp]: + "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b" +apply (cases "b=0", simp) +apply (simp add: nonzero_abs_divide) +done + +lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==> + abs x / y = abs (x / y)" + apply (subst abs_divide) + apply (simp add: order_less_imp_le) +done + +code_modulename SML + Fields Arith + +code_modulename OCaml + Fields Arith + +code_modulename Haskell + Fields Arith + +end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Groebner_Basis.thy Mon Feb 08 17:12:38 2010 +0100 @@ -621,7 +621,7 @@ val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, @{thm "divide_Numeral1"}, - @{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"}, + @{thm "Fields.divide_zero"}, @{thm "divide_Numeral0"}, @{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"}, @{thm "mult_num_frac"}, @{thm "mult_frac_num"}, @{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"}, diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Groups.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Groups.thy Mon Feb 08 17:12:38 2010 +0100 @@ -0,0 +1,1159 @@ +(* Title: HOL/Groups.thy + Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad +*) + +header {* Groups, also combined with orderings *} + +theory Groups +imports Lattices +uses "~~/src/Provers/Arith/abel_cancel.ML" +begin + +text {* + The theory of partially ordered groups is taken from the books: + \begin{itemize} + \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 + \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963 + \end{itemize} + Most of the used notions can also be looked up in + \begin{itemize} + \item \url{http://www.mathworld.com} by Eric Weisstein et. al. + \item \emph{Algebra I} by van der Waerden, Springer. + \end{itemize} +*} + +ML {* +structure Algebra_Simps = Named_Thms( + val name = "algebra_simps" + val description = "algebra simplification rules" +) +*} + +setup Algebra_Simps.setup + +text{* The rewrites accumulated in @{text algebra_simps} deal with the +classical algebraic structures of groups, rings and family. They simplify +terms by multiplying everything out (in case of a ring) and bringing sums and +products into a canonical form (by ordered rewriting). As a result it decides +group and ring equalities but also helps with inequalities. + +Of course it also works for fields, but it knows nothing about multiplicative +inverses or division. This is catered for by @{text field_simps}. *} + +subsection {* Semigroups and Monoids *} + +class semigroup_add = plus + + assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)" + +sublocale semigroup_add < plus!: semigroup plus proof +qed (fact add_assoc) + +class ab_semigroup_add = semigroup_add + + assumes add_commute [algebra_simps]: "a + b = b + a" + +sublocale ab_semigroup_add < plus!: abel_semigroup plus proof +qed (fact add_commute) + +context ab_semigroup_add +begin + +lemmas add_left_commute [algebra_simps] = plus.left_commute + +theorems add_ac = add_assoc add_commute add_left_commute + +end + +theorems add_ac = add_assoc add_commute add_left_commute + +class semigroup_mult = times + + assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)" + +sublocale semigroup_mult < times!: semigroup times proof +qed (fact mult_assoc) + +class ab_semigroup_mult = semigroup_mult + + assumes mult_commute [algebra_simps]: "a * b = b * a" + +sublocale ab_semigroup_mult < times!: abel_semigroup times proof +qed (fact mult_commute) + +context ab_semigroup_mult +begin + +lemmas mult_left_commute [algebra_simps] = times.left_commute + +theorems mult_ac = mult_assoc mult_commute mult_left_commute + +end + +theorems mult_ac = mult_assoc mult_commute mult_left_commute + +class ab_semigroup_idem_mult = ab_semigroup_mult + + assumes mult_idem: "x * x = x" + +sublocale ab_semigroup_idem_mult < times!: semilattice times proof +qed (fact mult_idem) + +context ab_semigroup_idem_mult +begin + +lemmas mult_left_idem = times.left_idem + +end + +class monoid_add = zero + semigroup_add + + assumes add_0_left [simp]: "0 + a = a" + and add_0_right [simp]: "a + 0 = a" + +lemma zero_reorient: "0 = x \ x = 0" +by (rule eq_commute) + +class comm_monoid_add = zero + ab_semigroup_add + + assumes add_0: "0 + a = a" +begin + +subclass monoid_add + proof qed (insert add_0, simp_all add: add_commute) + +end + +class monoid_mult = one + semigroup_mult + + assumes mult_1_left [simp]: "1 * a = a" + assumes mult_1_right [simp]: "a * 1 = a" + +lemma one_reorient: "1 = x \ x = 1" +by (rule eq_commute) + +class comm_monoid_mult = one + ab_semigroup_mult + + assumes mult_1: "1 * a = a" +begin + +subclass monoid_mult + proof qed (insert mult_1, simp_all add: mult_commute) + +end + +class cancel_semigroup_add = semigroup_add + + assumes add_left_imp_eq: "a + b = a + c \ b = c" + assumes add_right_imp_eq: "b + a = c + a \ b = c" +begin + +lemma add_left_cancel [simp]: + "a + b = a + c \ b = c" +by (blast dest: add_left_imp_eq) + +lemma add_right_cancel [simp]: + "b + a = c + a \ b = c" +by (blast dest: add_right_imp_eq) + +end + +class cancel_ab_semigroup_add = ab_semigroup_add + + assumes add_imp_eq: "a + b = a + c \ b = c" +begin + +subclass cancel_semigroup_add +proof + fix a b c :: 'a + assume "a + b = a + c" + then show "b = c" by (rule add_imp_eq) +next + fix a b c :: 'a + assume "b + a = c + a" + then have "a + b = a + c" by (simp only: add_commute) + then show "b = c" by (rule add_imp_eq) +qed + +end + +class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add + + +subsection {* Groups *} + +class group_add = minus + uminus + monoid_add + + assumes left_minus [simp]: "- a + a = 0" + assumes diff_minus: "a - b = a + (- b)" +begin + +lemma minus_unique: + assumes "a + b = 0" shows "- a = b" +proof - + have "- a = - a + (a + b)" using assms by simp + also have "\ = b" by (simp add: add_assoc [symmetric]) + finally show ?thesis . +qed + +lemmas equals_zero_I = minus_unique (* legacy name *) + +lemma minus_zero [simp]: "- 0 = 0" +proof - + have "0 + 0 = 0" by (rule add_0_right) + thus "- 0 = 0" by (rule minus_unique) +qed + +lemma minus_minus [simp]: "- (- a) = a" +proof - + have "- a + a = 0" by (rule left_minus) + thus "- (- a) = a" by (rule minus_unique) +qed + +lemma right_minus [simp]: "a + - a = 0" +proof - + have "a + - a = - (- a) + - a" by simp + also have "\ = 0" by (rule left_minus) + finally show ?thesis . +qed + +lemma minus_add_cancel: "- a + (a + b) = b" +by (simp add: add_assoc [symmetric]) + +lemma add_minus_cancel: "a + (- a + b) = b" +by (simp add: add_assoc [symmetric]) + +lemma minus_add: "- (a + b) = - b + - a" +proof - + have "(a + b) + (- b + - a) = 0" + by (simp add: add_assoc add_minus_cancel) + thus "- (a + b) = - b + - a" + by (rule minus_unique) +qed + +lemma right_minus_eq: "a - b = 0 \ a = b" +proof + assume "a - b = 0" + have "a = (a - b) + b" by (simp add:diff_minus add_assoc) + also have "\ = b" using `a - b = 0` by simp + finally show "a = b" . +next + assume "a = b" thus "a - b = 0" by (simp add: diff_minus) +qed + +lemma diff_self [simp]: "a - a = 0" +by (simp add: diff_minus) + +lemma diff_0 [simp]: "0 - a = - a" +by (simp add: diff_minus) + +lemma diff_0_right [simp]: "a - 0 = a" +by (simp add: diff_minus) + +lemma diff_minus_eq_add [simp]: "a - - b = a + b" +by (simp add: diff_minus) + +lemma neg_equal_iff_equal [simp]: + "- a = - b \ a = b" +proof + assume "- a = - b" + hence "- (- a) = - (- b)" by simp + thus "a = b" by simp +next + assume "a = b" + thus "- a = - b" by simp +qed + +lemma neg_equal_0_iff_equal [simp]: + "- a = 0 \ a = 0" +by (subst neg_equal_iff_equal [symmetric], simp) + +lemma neg_0_equal_iff_equal [simp]: + "0 = - a \ 0 = a" +by (subst neg_equal_iff_equal [symmetric], simp) + +text{*The next two equations can make the simplifier loop!*} + +lemma equation_minus_iff: + "a = - b \ b = - a" +proof - + have "- (- a) = - b \ - a = b" by (rule neg_equal_iff_equal) + thus ?thesis by (simp add: eq_commute) +qed + +lemma minus_equation_iff: + "- a = b \ - b = a" +proof - + have "- a = - (- b) \ a = -b" by (rule neg_equal_iff_equal) + thus ?thesis by (simp add: eq_commute) +qed + +lemma diff_add_cancel: "a - b + b = a" +by (simp add: diff_minus add_assoc) + +lemma add_diff_cancel: "a + b - b = a" +by (simp add: diff_minus add_assoc) + +declare diff_minus[symmetric, algebra_simps] + +lemma eq_neg_iff_add_eq_0: "a = - b \ a + b = 0" +proof + assume "a = - b" then show "a + b = 0" by simp +next + assume "a + b = 0" + moreover have "a + (b + - b) = (a + b) + - b" + by (simp only: add_assoc) + ultimately show "a = - b" by simp +qed + +end + +class ab_group_add = minus + uminus + comm_monoid_add + + assumes ab_left_minus: "- a + a = 0" + assumes ab_diff_minus: "a - b = a + (- b)" +begin + +subclass group_add + proof qed (simp_all add: ab_left_minus ab_diff_minus) + +subclass cancel_comm_monoid_add +proof + fix a b c :: 'a + assume "a + b = a + c" + then have "- a + a + b = - a + a + c" + unfolding add_assoc by simp + then show "b = c" by simp +qed + +lemma uminus_add_conv_diff[algebra_simps]: + "- a + b = b - a" +by (simp add:diff_minus add_commute) + +lemma minus_add_distrib [simp]: + "- (a + b) = - a + - b" +by (rule minus_unique) (simp add: add_ac) + +lemma minus_diff_eq [simp]: + "- (a - b) = b - a" +by (simp add: diff_minus add_commute) + +lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c" +by (simp add: diff_minus add_ac) + +lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b" +by (simp add: diff_minus add_ac) + +lemma diff_eq_eq[algebra_simps]: "a - b = c \ a = c + b" +by (auto simp add: diff_minus add_assoc) + +lemma eq_diff_eq[algebra_simps]: "a = c - b \ a + b = c" +by (auto simp add: diff_minus add_assoc) + +lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)" +by (simp add: diff_minus add_ac) + +lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b" +by (simp add: diff_minus add_ac) + +lemma eq_iff_diff_eq_0: "a = b \ a - b = 0" +by (simp add: algebra_simps) + +lemma diff_eq_0_iff_eq [simp, noatp]: "a - b = 0 \ a = b" +by (simp add: algebra_simps) + +end + +subsection {* (Partially) Ordered Groups *} + +class ordered_ab_semigroup_add = order + ab_semigroup_add + + assumes add_left_mono: "a \ b \ c + a \ c + b" +begin + +lemma add_right_mono: + "a \ b \ a + c \ b + c" +by (simp add: add_commute [of _ c] add_left_mono) + +text {* non-strict, in both arguments *} +lemma add_mono: + "a \ b \ c \ d \ a + c \ b + d" + apply (erule add_right_mono [THEN order_trans]) + apply (simp add: add_commute add_left_mono) + done + +end + +class ordered_cancel_ab_semigroup_add = + ordered_ab_semigroup_add + cancel_ab_semigroup_add +begin + +lemma add_strict_left_mono: + "a < b \ c + a < c + b" +by (auto simp add: less_le add_left_mono) + +lemma add_strict_right_mono: + "a < b \ a + c < b + c" +by (simp add: add_commute [of _ c] add_strict_left_mono) + +text{*Strict monotonicity in both arguments*} +lemma add_strict_mono: + "a < b \ c < d \ a + c < b + d" +apply (erule add_strict_right_mono [THEN less_trans]) +apply (erule add_strict_left_mono) +done + +lemma add_less_le_mono: + "a < b \ c \ d \ a + c < b + d" +apply (erule add_strict_right_mono [THEN less_le_trans]) +apply (erule add_left_mono) +done + +lemma add_le_less_mono: + "a \ b \ c < d \ a + c < b + d" +apply (erule add_right_mono [THEN le_less_trans]) +apply (erule add_strict_left_mono) +done + +end + +class ordered_ab_semigroup_add_imp_le = + ordered_cancel_ab_semigroup_add + + assumes add_le_imp_le_left: "c + a \ c + b \ a \ b" +begin + +lemma add_less_imp_less_left: + assumes less: "c + a < c + b" shows "a < b" +proof - + from less have le: "c + a <= c + b" by (simp add: order_le_less) + have "a <= b" + apply (insert le) + apply (drule add_le_imp_le_left) + by (insert le, drule add_le_imp_le_left, assumption) + moreover have "a \ b" + proof (rule ccontr) + assume "~(a \ b)" + then have "a = b" by simp + then have "c + a = c + b" by simp + with less show "False"by simp + qed + ultimately show "a < b" by (simp add: order_le_less) +qed + +lemma add_less_imp_less_right: + "a + c < b + c \ a < b" +apply (rule add_less_imp_less_left [of c]) +apply (simp add: add_commute) +done + +lemma add_less_cancel_left [simp]: + "c + a < c + b \ a < b" +by (blast intro: add_less_imp_less_left add_strict_left_mono) + +lemma add_less_cancel_right [simp]: + "a + c < b + c \ a < b" +by (blast intro: add_less_imp_less_right add_strict_right_mono) + +lemma add_le_cancel_left [simp]: + "c + a \ c + b \ a \ b" +by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) + +lemma add_le_cancel_right [simp]: + "a + c \ b + c \ a \ b" +by (simp add: add_commute [of a c] add_commute [of b c]) + +lemma add_le_imp_le_right: + "a + c \ b + c \ a \ b" +by simp + +lemma max_add_distrib_left: + "max x y + z = max (x + z) (y + z)" + unfolding max_def by auto + +lemma min_add_distrib_left: + "min x y + z = min (x + z) (y + z)" + unfolding min_def by auto + +end + +subsection {* Support for reasoning about signs *} + +class ordered_comm_monoid_add = + ordered_cancel_ab_semigroup_add + comm_monoid_add +begin + +lemma add_pos_nonneg: + assumes "0 < a" and "0 \ b" shows "0 < a + b" +proof - + have "0 + 0 < a + b" + using assms by (rule add_less_le_mono) + then show ?thesis by simp +qed + +lemma add_pos_pos: + assumes "0 < a" and "0 < b" shows "0 < a + b" +by (rule add_pos_nonneg) (insert assms, auto) + +lemma add_nonneg_pos: + assumes "0 \ a" and "0 < b" shows "0 < a + b" +proof - + have "0 + 0 < a + b" + using assms by (rule add_le_less_mono) + then show ?thesis by simp +qed + +lemma add_nonneg_nonneg: + assumes "0 \ a" and "0 \ b" shows "0 \ a + b" +proof - + have "0 + 0 \ a + b" + using assms by (rule add_mono) + then show ?thesis by simp +qed + +lemma add_neg_nonpos: + assumes "a < 0" and "b \ 0" shows "a + b < 0" +proof - + have "a + b < 0 + 0" + using assms by (rule add_less_le_mono) + then show ?thesis by simp +qed + +lemma add_neg_neg: + assumes "a < 0" and "b < 0" shows "a + b < 0" +by (rule add_neg_nonpos) (insert assms, auto) + +lemma add_nonpos_neg: + assumes "a \ 0" and "b < 0" shows "a + b < 0" +proof - + have "a + b < 0 + 0" + using assms by (rule add_le_less_mono) + then show ?thesis by simp +qed + +lemma add_nonpos_nonpos: + assumes "a \ 0" and "b \ 0" shows "a + b \ 0" +proof - + have "a + b \ 0 + 0" + using assms by (rule add_mono) + then show ?thesis by simp +qed + +lemmas add_sign_intros = + add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg + add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos + +lemma add_nonneg_eq_0_iff: + assumes x: "0 \ x" and y: "0 \ y" + shows "x + y = 0 \ x = 0 \ y = 0" +proof (intro iffI conjI) + have "x = x + 0" by simp + also have "x + 0 \ x + y" using y by (rule add_left_mono) + also assume "x + y = 0" + also have "0 \ x" using x . + finally show "x = 0" . +next + have "y = 0 + y" by simp + also have "0 + y \ x + y" using x by (rule add_right_mono) + also assume "x + y = 0" + also have "0 \ y" using y . + finally show "y = 0" . +next + assume "x = 0 \ y = 0" + then show "x + y = 0" by simp +qed + +end + +class ordered_ab_group_add = + ab_group_add + ordered_ab_semigroup_add +begin + +subclass ordered_cancel_ab_semigroup_add .. + +subclass ordered_ab_semigroup_add_imp_le +proof + fix a b c :: 'a + assume "c + a \ c + b" + hence "(-c) + (c + a) \ (-c) + (c + b)" by (rule add_left_mono) + hence "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add_assoc) + thus "a \ b" by simp +qed + +subclass ordered_comm_monoid_add .. + +lemma max_diff_distrib_left: + shows "max x y - z = max (x - z) (y - z)" +by (simp add: diff_minus, rule max_add_distrib_left) + +lemma min_diff_distrib_left: + shows "min x y - z = min (x - z) (y - z)" +by (simp add: diff_minus, rule min_add_distrib_left) + +lemma le_imp_neg_le: + assumes "a \ b" shows "-b \ -a" +proof - + have "-a+a \ -a+b" using `a \ b` by (rule add_left_mono) + hence "0 \ -a+b" by simp + hence "0 + (-b) \ (-a + b) + (-b)" by (rule add_right_mono) + thus ?thesis by (simp add: add_assoc) +qed + +lemma neg_le_iff_le [simp]: "- b \ - a \ a \ b" +proof + assume "- b \ - a" + hence "- (- a) \ - (- b)" by (rule le_imp_neg_le) + thus "a\b" by simp +next + assume "a\b" + thus "-b \ -a" by (rule le_imp_neg_le) +qed + +lemma neg_le_0_iff_le [simp]: "- a \ 0 \ 0 \ a" +by (subst neg_le_iff_le [symmetric], simp) + +lemma neg_0_le_iff_le [simp]: "0 \ - a \ a \ 0" +by (subst neg_le_iff_le [symmetric], simp) + +lemma neg_less_iff_less [simp]: "- b < - a \ a < b" +by (force simp add: less_le) + +lemma neg_less_0_iff_less [simp]: "- a < 0 \ 0 < a" +by (subst neg_less_iff_less [symmetric], simp) + +lemma neg_0_less_iff_less [simp]: "0 < - a \ a < 0" +by (subst neg_less_iff_less [symmetric], simp) + +text{*The next several equations can make the simplifier loop!*} + +lemma less_minus_iff: "a < - b \ b < - a" +proof - + have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less) + thus ?thesis by simp +qed + +lemma minus_less_iff: "- a < b \ - b < a" +proof - + have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less) + thus ?thesis by simp +qed + +lemma le_minus_iff: "a \ - b \ b \ - a" +proof - + have mm: "!! a (b::'a). (-(-a)) < -b \ -(-b) < -a" by (simp only: minus_less_iff) + have "(- (- a) <= -b) = (b <= - a)" + apply (auto simp only: le_less) + apply (drule mm) + apply (simp_all) + apply (drule mm[simplified], assumption) + done + then show ?thesis by simp +qed + +lemma minus_le_iff: "- a \ b \ - b \ a" +by (auto simp add: le_less minus_less_iff) + +lemma less_iff_diff_less_0: "a < b \ a - b < 0" +proof - + have "(a < b) = (a + (- b) < b + (-b))" + by (simp only: add_less_cancel_right) + also have "... = (a - b < 0)" by (simp add: diff_minus) + finally show ?thesis . +qed + +lemma diff_less_eq[algebra_simps]: "a - b < c \ a < c + b" +apply (subst less_iff_diff_less_0 [of a]) +apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst]) +apply (simp add: diff_minus add_ac) +done + +lemma less_diff_eq[algebra_simps]: "a < c - b \ a + b < c" +apply (subst less_iff_diff_less_0 [of "plus a b"]) +apply (subst less_iff_diff_less_0 [of a]) +apply (simp add: diff_minus add_ac) +done + +lemma diff_le_eq[algebra_simps]: "a - b \ c \ a \ c + b" +by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel) + +lemma le_diff_eq[algebra_simps]: "a \ c - b \ a + b \ c" +by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel) + +lemma le_iff_diff_le_0: "a \ b \ a - b \ 0" +by (simp add: algebra_simps) + +text{*Legacy - use @{text algebra_simps} *} +lemmas group_simps[noatp] = algebra_simps + +end + +text{*Legacy - use @{text algebra_simps} *} +lemmas group_simps[noatp] = algebra_simps + +class linordered_ab_semigroup_add = + linorder + ordered_ab_semigroup_add + +class linordered_cancel_ab_semigroup_add = + linorder + ordered_cancel_ab_semigroup_add +begin + +subclass linordered_ab_semigroup_add .. + +subclass ordered_ab_semigroup_add_imp_le +proof + fix a b c :: 'a + assume le: "c + a <= c + b" + show "a <= b" + proof (rule ccontr) + assume w: "~ a \ b" + hence "b <= a" by (simp add: linorder_not_le) + hence le2: "c + b <= c + a" by (rule add_left_mono) + have "a = b" + apply (insert le) + apply (insert le2) + apply (drule antisym, simp_all) + done + with w show False + by (simp add: linorder_not_le [symmetric]) + qed +qed + +end + +class linordered_ab_group_add = linorder + ordered_ab_group_add +begin + +subclass linordered_cancel_ab_semigroup_add .. + +lemma neg_less_eq_nonneg [simp]: + "- a \ a \ 0 \ a" +proof + assume A: "- a \ a" show "0 \ a" + proof (rule classical) + assume "\ 0 \ a" + then have "a < 0" by auto + with A have "- a < 0" by (rule le_less_trans) + then show ?thesis by auto + qed +next + assume A: "0 \ a" show "- a \ a" + proof (rule order_trans) + show "- a \ 0" using A by (simp add: minus_le_iff) + next + show "0 \ a" using A . + qed +qed + +lemma neg_less_nonneg [simp]: + "- a < a \ 0 < a" +proof + assume A: "- a < a" show "0 < a" + proof (rule classical) + assume "\ 0 < a" + then have "a \ 0" by auto + with A have "- a < 0" by (rule less_le_trans) + then show ?thesis by auto + qed +next + assume A: "0 < a" show "- a < a" + proof (rule less_trans) + show "- a < 0" using A by (simp add: minus_le_iff) + next + show "0 < a" using A . + qed +qed + +lemma less_eq_neg_nonpos [simp]: + "a \ - a \ a \ 0" +proof + assume A: "a \ - a" show "a \ 0" + proof (rule classical) + assume "\ a \ 0" + then have "0 < a" by auto + then have "0 < - a" using A by (rule less_le_trans) + then show ?thesis by auto + qed +next + assume A: "a \ 0" show "a \ - a" + proof (rule order_trans) + show "0 \ - a" using A by (simp add: minus_le_iff) + next + show "a \ 0" using A . + qed +qed + +lemma equal_neg_zero [simp]: + "a = - a \ a = 0" +proof + assume "a = 0" then show "a = - a" by simp +next + assume A: "a = - a" show "a = 0" + proof (cases "0 \ a") + case True with A have "0 \ - a" by auto + with le_minus_iff have "a \ 0" by simp + with True show ?thesis by (auto intro: order_trans) + next + case False then have B: "a \ 0" by auto + with A have "- a \ 0" by auto + with B show ?thesis by (auto intro: order_trans) + qed +qed + +lemma neg_equal_zero [simp]: + "- a = a \ a = 0" + by (auto dest: sym) + +lemma double_zero [simp]: + "a + a = 0 \ a = 0" +proof + assume assm: "a + a = 0" + then have a: "- a = a" by (rule minus_unique) + then show "a = 0" by (simp add: neg_equal_zero) +qed simp + +lemma double_zero_sym [simp]: + "0 = a + a \ a = 0" + by (rule, drule sym) simp_all + +lemma zero_less_double_add_iff_zero_less_single_add [simp]: + "0 < a + a \ 0 < a" +proof + assume "0 < a + a" + then have "0 - a < a" by (simp only: diff_less_eq) + then have "- a < a" by simp + then show "0 < a" by (simp add: neg_less_nonneg) +next + assume "0 < a" + with this have "0 + 0 < a + a" + by (rule add_strict_mono) + then show "0 < a + a" by simp +qed + +lemma zero_le_double_add_iff_zero_le_single_add [simp]: + "0 \ a + a \ 0 \ a" + by (auto simp add: le_less) + +lemma double_add_less_zero_iff_single_add_less_zero [simp]: + "a + a < 0 \ a < 0" +proof - + have "\ a + a < 0 \ \ a < 0" + by (simp add: not_less) + then show ?thesis by simp +qed + +lemma double_add_le_zero_iff_single_add_le_zero [simp]: + "a + a \ 0 \ a \ 0" +proof - + have "\ a + a \ 0 \ \ a \ 0" + by (simp add: not_le) + then show ?thesis by simp +qed + +lemma le_minus_self_iff: + "a \ - a \ a \ 0" +proof - + from add_le_cancel_left [of "- a" "a + a" 0] + have "a \ - a \ a + a \ 0" + by (simp add: add_assoc [symmetric]) + thus ?thesis by simp +qed + +lemma minus_le_self_iff: + "- a \ a \ 0 \ a" +proof - + from add_le_cancel_left [of "- a" 0 "a + a"] + have "- a \ a \ 0 \ a + a" + by (simp add: add_assoc [symmetric]) + thus ?thesis by simp +qed + +lemma minus_max_eq_min: + "- max x y = min (-x) (-y)" + by (auto simp add: max_def min_def) + +lemma minus_min_eq_max: + "- min x y = max (-x) (-y)" + by (auto simp add: max_def min_def) + +end + +-- {* FIXME localize the following *} + +lemma add_increasing: + fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" + shows "[|0\a; b\c|] ==> b \ a + c" +by (insert add_mono [of 0 a b c], simp) + +lemma add_increasing2: + fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" + shows "[|0\c; b\a|] ==> b \ a + c" +by (simp add:add_increasing add_commute[of a]) + +lemma add_strict_increasing: + fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" + shows "[|0c|] ==> b < a + c" +by (insert add_less_le_mono [of 0 a b c], simp) + +lemma add_strict_increasing2: + fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" + shows "[|0\a; b b < a + c" +by (insert add_le_less_mono [of 0 a b c], simp) + + +class ordered_ab_group_add_abs = ordered_ab_group_add + abs + + assumes abs_ge_zero [simp]: "\a\ \ 0" + and abs_ge_self: "a \ \a\" + and abs_leI: "a \ b \ - a \ b \ \a\ \ b" + and abs_minus_cancel [simp]: "\-a\ = \a\" + and abs_triangle_ineq: "\a + b\ \ \a\ + \b\" +begin + +lemma abs_minus_le_zero: "- \a\ \ 0" + unfolding neg_le_0_iff_le by simp + +lemma abs_of_nonneg [simp]: + assumes nonneg: "0 \ a" shows "\a\ = a" +proof (rule antisym) + from nonneg le_imp_neg_le have "- a \ 0" by simp + from this nonneg have "- a \ a" by (rule order_trans) + then show "\a\ \ a" by (auto intro: abs_leI) +qed (rule abs_ge_self) + +lemma abs_idempotent [simp]: "\\a\\ = \a\" +by (rule antisym) + (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"]) + +lemma abs_eq_0 [simp]: "\a\ = 0 \ a = 0" +proof - + have "\a\ = 0 \ a = 0" + proof (rule antisym) + assume zero: "\a\ = 0" + with abs_ge_self show "a \ 0" by auto + from zero have "\-a\ = 0" by simp + with abs_ge_self [of "uminus a"] have "- a \ 0" by auto + with neg_le_0_iff_le show "0 \ a" by auto + qed + then show ?thesis by auto +qed + +lemma abs_zero [simp]: "\0\ = 0" +by simp + +lemma abs_0_eq [simp, noatp]: "0 = \a\ \ a = 0" +proof - + have "0 = \a\ \ \a\ = 0" by (simp only: eq_ac) + thus ?thesis by simp +qed + +lemma abs_le_zero_iff [simp]: "\a\ \ 0 \ a = 0" +proof + assume "\a\ \ 0" + then have "\a\ = 0" by (rule antisym) simp + thus "a = 0" by simp +next + assume "a = 0" + thus "\a\ \ 0" by simp +qed + +lemma zero_less_abs_iff [simp]: "0 < \a\ \ a \ 0" +by (simp add: less_le) + +lemma abs_not_less_zero [simp]: "\ \a\ < 0" +proof - + have a: "\x y. x \ y \ \ y < x" by auto + show ?thesis by (simp add: a) +qed + +lemma abs_ge_minus_self: "- a \ \a\" +proof - + have "- a \ \-a\" by (rule abs_ge_self) + then show ?thesis by simp +qed + +lemma abs_minus_commute: + "\a - b\ = \b - a\" +proof - + have "\a - b\ = \- (a - b)\" by (simp only: abs_minus_cancel) + also have "... = \b - a\" by simp + finally show ?thesis . +qed + +lemma abs_of_pos: "0 < a \ \a\ = a" +by (rule abs_of_nonneg, rule less_imp_le) + +lemma abs_of_nonpos [simp]: + assumes "a \ 0" shows "\a\ = - a" +proof - + let ?b = "- a" + have "- ?b \ 0 \ \- ?b\ = - (- ?b)" + unfolding abs_minus_cancel [of "?b"] + unfolding neg_le_0_iff_le [of "?b"] + unfolding minus_minus by (erule abs_of_nonneg) + then show ?thesis using assms by auto +qed + +lemma abs_of_neg: "a < 0 \ \a\ = - a" +by (rule abs_of_nonpos, rule less_imp_le) + +lemma abs_le_D1: "\a\ \ b \ a \ b" +by (insert abs_ge_self, blast intro: order_trans) + +lemma abs_le_D2: "\a\ \ b \ - a \ b" +by (insert abs_le_D1 [of "uminus a"], simp) + +lemma abs_le_iff: "\a\ \ b \ a \ b \ - a \ b" +by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2) + +lemma abs_triangle_ineq2: "\a\ - \b\ \ \a - b\" + apply (simp add: algebra_simps) + apply (subgoal_tac "abs a = abs (plus b (minus a b))") + apply (erule ssubst) + apply (rule abs_triangle_ineq) + apply (rule arg_cong[of _ _ abs]) + apply (simp add: algebra_simps) +done + +lemma abs_triangle_ineq3: "\\a\ - \b\\ \ \a - b\" + apply (subst abs_le_iff) + apply auto + apply (rule abs_triangle_ineq2) + apply (subst abs_minus_commute) + apply (rule abs_triangle_ineq2) +done + +lemma abs_triangle_ineq4: "\a - b\ \ \a\ + \b\" +proof - + have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl) + also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq) + finally show ?thesis by simp +qed + +lemma abs_diff_triangle_ineq: "\a + b - (c + d)\ \ \a - c\ + \b - d\" +proof - + have "\a + b - (c+d)\ = \(a-c) + (b-d)\" by (simp add: diff_minus add_ac) + also have "... \ \a-c\ + \b-d\" by (rule abs_triangle_ineq) + finally show ?thesis . +qed + +lemma abs_add_abs [simp]: + "\\a\ + \b\\ = \a\ + \b\" (is "?L = ?R") +proof (rule antisym) + show "?L \ ?R" by(rule abs_ge_self) +next + have "?L \ \\a\\ + \\b\\" by(rule abs_triangle_ineq) + also have "\ = ?R" by simp + finally show "?L \ ?R" . +qed + +end + +text {* Needed for abelian cancellation simprocs: *} + +lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)" +apply (subst add_left_commute) +apply (subst add_left_cancel) +apply simp +done + +lemma add_cancel_end: "(x + (y + z) = y) = (x = - (z::'a::ab_group_add))" +apply (subst add_cancel_21[of _ _ _ 0, simplified]) +apply (simp add: add_right_cancel[symmetric, of "x" "-z" "z", simplified]) +done + +lemma less_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \ (x < y) = (x' < y')" +by (simp add: less_iff_diff_less_0[of x y] less_iff_diff_less_0[of x' y']) + +lemma le_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \ (y <= x) = (y' <= x')" +apply (simp add: le_iff_diff_le_0[of y x] le_iff_diff_le_0[of y' x']) +apply (simp add: neg_le_iff_le[symmetric, of "y-x" 0] neg_le_iff_le[symmetric, of "y'-x'" 0]) +done + +lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \ (x = y) = (x' = y')" +by (simp only: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y']) + +lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)" +by (simp add: diff_minus) + +lemma le_add_right_mono: + assumes + "a <= b + (c::'a::ordered_ab_group_add)" + "c <= d" + shows "a <= b + d" + apply (rule_tac order_trans[where y = "b+c"]) + apply (simp_all add: prems) + done + + +subsection {* Tools setup *} + +lemma add_mono_thms_linordered_semiring [noatp]: + fixes i j k :: "'a\ordered_ab_semigroup_add" + shows "i \ j \ k \ l \ i + k \ j + l" + and "i = j \ k \ l \ i + k \ j + l" + and "i \ j \ k = l \ i + k \ j + l" + and "i = j \ k = l \ i + k = j + l" +by (rule add_mono, clarify+)+ + +lemma add_mono_thms_linordered_field [noatp]: + fixes i j k :: "'a\ordered_cancel_ab_semigroup_add" + shows "i < j \ k = l \ i + k < j + l" + and "i = j \ k < l \ i + k < j + l" + and "i < j \ k \ l \ i + k < j + l" + and "i \ j \ k < l \ i + k < j + l" + and "i < j \ k < l \ i + k < j + l" +by (auto intro: add_strict_right_mono add_strict_left_mono + add_less_le_mono add_le_less_mono add_strict_mono) + +text{*Simplification of @{term "x-y < 0"}, etc.*} +lemmas diff_less_0_iff_less [simp, noatp] = less_iff_diff_less_0 [symmetric] +lemmas diff_le_0_iff_le [simp, noatp] = le_iff_diff_le_0 [symmetric] + +ML {* +structure ab_group_add_cancel = Abel_Cancel +( + +(* term order for abelian groups *) + +fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a') + [@{const_name Algebras.zero}, @{const_name Algebras.plus}, + @{const_name Algebras.uminus}, @{const_name Algebras.minus}] + | agrp_ord _ = ~1; + +fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS); + +local + val ac1 = mk_meta_eq @{thm add_assoc}; + val ac2 = mk_meta_eq @{thm add_commute}; + val ac3 = mk_meta_eq @{thm add_left_commute}; + fun solve_add_ac thy _ (_ $ (Const (@{const_name Algebras.plus},_) $ _ $ _) $ _) = + SOME ac1 + | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Algebras.plus},_) $ y $ z)) = + if termless_agrp (y, x) then SOME ac3 else NONE + | solve_add_ac thy _ (_ $ x $ y) = + if termless_agrp (y, x) then SOME ac2 else NONE + | solve_add_ac thy _ _ = NONE +in + val add_ac_proc = Simplifier.simproc @{theory} + "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac; +end; + +val eq_reflection = @{thm eq_reflection}; + +val T = @{typ "'a::ab_group_add"}; + +val cancel_ss = HOL_basic_ss settermless termless_agrp + addsimprocs [add_ac_proc] addsimps + [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_def}, + @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero}, + @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel}, + @{thm minus_add_cancel}]; + +val sum_pats = [@{cterm "x + y::'a::ab_group_add"}, @{cterm "x - y::'a::ab_group_add"}]; + +val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}]; + +val dest_eqI = + fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of; + +); +*} + +ML {* + Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv]; +*} + +code_modulename SML + Groups Arith + +code_modulename OCaml + Groups Arith + +code_modulename Haskell + Groups Arith + +end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Import/HOL/arithmetic.imp --- a/src/HOL/Import/HOL/arithmetic.imp Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Import/HOL/arithmetic.imp Mon Feb 08 17:12:38 2010 +0100 @@ -162,12 +162,12 @@ "LESS_OR" > "Nat.Suc_leI" "LESS_NOT_SUC" > "HOL4Base.arithmetic.LESS_NOT_SUC" "LESS_MULT_MONO" > "Nat.Suc_mult_less_cancel1" - "LESS_MULT2" > "Ring_and_Field.mult_pos_pos" + "LESS_MULT2" > "Rings.mult_pos_pos" "LESS_MONO_REV" > "Nat.Suc_less_SucD" "LESS_MONO_MULT" > "Nat.mult_le_mono1" "LESS_MONO_EQ" > "Nat.Suc_less_eq" - "LESS_MONO_ADD_INV" > "OrderedGroup.add_less_imp_less_right" - "LESS_MONO_ADD_EQ" > "OrderedGroup.add_less_cancel_right" + "LESS_MONO_ADD_INV" > "Groups.add_less_imp_less_right" + "LESS_MONO_ADD_EQ" > "Groups.add_less_cancel_right" "LESS_MONO_ADD" > "Nat.add_less_mono1" "LESS_MOD" > "Divides.mod_less" "LESS_LESS_SUC" > "HOL4Base.arithmetic.LESS_LESS_SUC" @@ -180,7 +180,7 @@ "LESS_EQ_SUC_REFL" > "HOL4Base.arithmetic.LESS_EQ_SUC_REFL" "LESS_EQ_SUB_LESS" > "HOL4Base.arithmetic.LESS_EQ_SUB_LESS" "LESS_EQ_REFL" > "Finite_Set.max.f_below.below_refl" - "LESS_EQ_MONO_ADD_EQ" > "OrderedGroup.add_le_cancel_right" + "LESS_EQ_MONO_ADD_EQ" > "Groups.add_le_cancel_right" "LESS_EQ_MONO" > "Nat.Suc_le_mono" "LESS_EQ_LESS_TRANS" > "Nat.le_less_trans" "LESS_EQ_LESS_EQ_MONO" > "Nat.add_le_mono" diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Import/HOL/divides.imp --- a/src/HOL/Import/HOL/divides.imp Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Import/HOL/divides.imp Mon Feb 08 17:12:38 2010 +0100 @@ -9,16 +9,16 @@ "divides_def" > "HOL4Compat.divides_def" "ONE_DIVIDES_ALL" > "HOL4Base.divides.ONE_DIVIDES_ALL" "NOT_LT_DIV" > "NatSimprocs.nat_dvd_not_less" - "DIVIDES_TRANS" > "Ring_and_Field.dvd_trans" - "DIVIDES_SUB" > "Ring_and_Field.dvd_diff" - "DIVIDES_REFL" > "Ring_and_Field.dvd_refl" + "DIVIDES_TRANS" > "Rings.dvd_trans" + "DIVIDES_SUB" > "Rings.dvd_diff" + "DIVIDES_REFL" > "Rings.dvd_refl" "DIVIDES_MULT_LEFT" > "HOL4Base.divides.DIVIDES_MULT_LEFT" "DIVIDES_MULT" > "Divides.dvd_mult2" "DIVIDES_LE" > "Divides.dvd_imp_le" "DIVIDES_FACT" > "HOL4Base.divides.DIVIDES_FACT" "DIVIDES_ANTISYM" > "Divides.dvd_antisym" "DIVIDES_ADD_2" > "HOL4Base.divides.DIVIDES_ADD_2" - "DIVIDES_ADD_1" > "Ring_and_Field.dvd_add" + "DIVIDES_ADD_1" > "Rings.dvd_add" "ALL_DIVIDES_0" > "Divides.dvd_0_right" end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Import/HOL/prob_extra.imp --- a/src/HOL/Import/HOL/prob_extra.imp Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Import/HOL/prob_extra.imp Mon Feb 08 17:12:38 2010 +0100 @@ -23,9 +23,9 @@ "REAL_SUP_LE_X" > "HOL4Prob.prob_extra.REAL_SUP_LE_X" "REAL_SUP_EXISTS_UNIQUE" > "HOL4Prob.prob_extra.REAL_SUP_EXISTS_UNIQUE" "REAL_POW" > "RealPow.realpow_real_of_nat" - "REAL_LE_INV_LE" > "Ring_and_Field.le_imp_inverse_le" + "REAL_LE_INV_LE" > "Rings.le_imp_inverse_le" "REAL_LE_EQ" > "Set.basic_trans_rules_26" - "REAL_INVINV_ALL" > "Ring_and_Field.inverse_inverse_eq" + "REAL_INVINV_ALL" > "Rings.inverse_inverse_eq" "REAL_INF_MIN" > "HOL4Prob.prob_extra.REAL_INF_MIN" "RAND_THM" > "HOL.arg_cong" "POW_HALF_TWICE" > "HOL4Prob.prob_extra.POW_HALF_TWICE" diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Import/HOL/real.imp --- a/src/HOL/Import/HOL/real.imp Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Import/HOL/real.imp Mon Feb 08 17:12:38 2010 +0100 @@ -25,13 +25,13 @@ "sumc" > "HOL4Real.real.sumc" "sum_def" > "HOL4Real.real.sum_def" "sum" > "HOL4Real.real.sum" - "real_sub" > "OrderedGroup.diff_def" + "real_sub" > "Groups.diff_def" "real_of_num" > "HOL4Compat.real_of_num" "real_lte" > "HOL4Compat.real_lte" "real_lt" > "Orderings.linorder_not_le" "real_gt" > "HOL4Compat.real_gt" "real_ge" > "HOL4Compat.real_ge" - "real_div" > "Ring_and_Field.field_class.divide_inverse" + "real_div" > "Rings.field_class.divide_inverse" "pow" > "HOL4Compat.pow" "abs" > "HOL4Compat.abs" "SUP_LEMMA3" > "HOL4Real.real.SUP_LEMMA3" @@ -74,24 +74,24 @@ "REAL_SUB_TRIANGLE" > "HOL4Real.real.REAL_SUB_TRIANGLE" "REAL_SUB_SUB2" > "HOL4Real.real.REAL_SUB_SUB2" "REAL_SUB_SUB" > "HOL4Real.real.REAL_SUB_SUB" - "REAL_SUB_RZERO" > "OrderedGroup.diff_0_right" - "REAL_SUB_RNEG" > "OrderedGroup.diff_minus_eq_add" - "REAL_SUB_REFL" > "OrderedGroup.diff_self" - "REAL_SUB_RDISTRIB" > "Ring_and_Field.ring_eq_simps_3" + "REAL_SUB_RZERO" > "Groups.diff_0_right" + "REAL_SUB_RNEG" > "Groups.diff_minus_eq_add" + "REAL_SUB_REFL" > "Groups.diff_self" + "REAL_SUB_RDISTRIB" > "Rings.ring_eq_simps_3" "REAL_SUB_NEG2" > "HOL4Real.real.REAL_SUB_NEG2" - "REAL_SUB_LZERO" > "OrderedGroup.diff_0" + "REAL_SUB_LZERO" > "Groups.diff_0" "REAL_SUB_LT" > "HOL4Real.real.REAL_SUB_LT" "REAL_SUB_LNEG" > "HOL4Real.real.REAL_SUB_LNEG" "REAL_SUB_LE" > "HOL4Real.real.REAL_SUB_LE" - "REAL_SUB_LDISTRIB" > "Ring_and_Field.ring_eq_simps_4" + "REAL_SUB_LDISTRIB" > "Rings.ring_eq_simps_4" "REAL_SUB_INV2" > "HOL4Real.real.REAL_SUB_INV2" "REAL_SUB_ADD2" > "HOL4Real.real.REAL_SUB_ADD2" - "REAL_SUB_ADD" > "OrderedGroup.diff_add_cancel" - "REAL_SUB_ABS" > "OrderedGroup.abs_triangle_ineq2" - "REAL_SUB_0" > "OrderedGroup.diff_eq_0_iff_eq" + "REAL_SUB_ADD" > "Groups.diff_add_cancel" + "REAL_SUB_ABS" > "Groups.abs_triangle_ineq2" + "REAL_SUB_0" > "Groups.diff_eq_0_iff_eq" "REAL_RNEG_UNIQ" > "RealDef.real_add_eq_0_iff" - "REAL_RINV_UNIQ" > "Ring_and_Field.inverse_unique" - "REAL_RDISTRIB" > "Ring_and_Field.ring_eq_simps_1" + "REAL_RINV_UNIQ" > "Rings.inverse_unique" + "REAL_RDISTRIB" > "Rings.ring_eq_simps_1" "REAL_POW_POW" > "Power.power_mult" "REAL_POW_MONO_LT" > "HOL4Real.real.REAL_POW_MONO_LT" "REAL_POW_LT2" > "HOL4Real.real.REAL_POW_LT2" @@ -103,7 +103,7 @@ "REAL_POS_NZ" > "HOL4Real.real.REAL_POS_NZ" "REAL_POS" > "RealDef.real_of_nat_ge_zero" "REAL_POASQ" > "HOL4Real.real.REAL_POASQ" - "REAL_OVER1" > "Ring_and_Field.divide_1" + "REAL_OVER1" > "Rings.divide_1" "REAL_OF_NUM_SUC" > "RealDef.real_of_nat_Suc" "REAL_OF_NUM_POW" > "RealPow.realpow_real_of_nat" "REAL_OF_NUM_MUL" > "RealDef.real_of_nat_mult" @@ -113,172 +113,172 @@ "REAL_NZ_IMP_LT" > "HOL4Real.real.REAL_NZ_IMP_LT" "REAL_NOT_LT" > "HOL4Compat.real_lte" "REAL_NOT_LE" > "Orderings.linorder_not_le" - "REAL_NEG_SUB" > "OrderedGroup.minus_diff_eq" - "REAL_NEG_RMUL" > "Ring_and_Field.mult_minus_right" - "REAL_NEG_NEG" > "OrderedGroup.minus_minus" - "REAL_NEG_MUL2" > "Ring_and_Field.minus_mult_minus" + "REAL_NEG_SUB" > "Groups.minus_diff_eq" + "REAL_NEG_RMUL" > "Rings.mult_minus_right" + "REAL_NEG_NEG" > "Groups.minus_minus" + "REAL_NEG_MUL2" > "Rings.minus_mult_minus" "REAL_NEG_MINUS1" > "HOL4Real.real.REAL_NEG_MINUS1" - "REAL_NEG_LT0" > "OrderedGroup.neg_less_0_iff_less" - "REAL_NEG_LMUL" > "Ring_and_Field.mult_minus_left" - "REAL_NEG_LE0" > "OrderedGroup.neg_le_0_iff_le" - "REAL_NEG_INV" > "Ring_and_Field.nonzero_inverse_minus_eq" - "REAL_NEG_GT0" > "OrderedGroup.neg_0_less_iff_less" - "REAL_NEG_GE0" > "OrderedGroup.neg_0_le_iff_le" - "REAL_NEG_EQ0" > "OrderedGroup.neg_equal_0_iff_equal" + "REAL_NEG_LT0" > "Groups.neg_less_0_iff_less" + "REAL_NEG_LMUL" > "Rings.mult_minus_left" + "REAL_NEG_LE0" > "Groups.neg_le_0_iff_le" + "REAL_NEG_INV" > "Rings.nonzero_inverse_minus_eq" + "REAL_NEG_GT0" > "Groups.neg_0_less_iff_less" + "REAL_NEG_GE0" > "Groups.neg_0_le_iff_le" + "REAL_NEG_EQ0" > "Groups.neg_equal_0_iff_equal" "REAL_NEG_EQ" > "HOL4Real.real.REAL_NEG_EQ" - "REAL_NEG_ADD" > "OrderedGroup.minus_add_distrib" - "REAL_NEG_0" > "OrderedGroup.minus_zero" - "REAL_NEGNEG" > "OrderedGroup.minus_minus" + "REAL_NEG_ADD" > "Groups.minus_add_distrib" + "REAL_NEG_0" > "Groups.minus_zero" + "REAL_NEGNEG" > "Groups.minus_minus" "REAL_MUL_SYM" > "Int.zmult_ac_2" - "REAL_MUL_RZERO" > "Ring_and_Field.mult_zero_right" - "REAL_MUL_RNEG" > "Ring_and_Field.mult_minus_right" - "REAL_MUL_RINV" > "Ring_and_Field.right_inverse" + "REAL_MUL_RZERO" > "Rings.mult_zero_right" + "REAL_MUL_RNEG" > "Rings.mult_minus_right" + "REAL_MUL_RINV" > "Rings.right_inverse" "REAL_MUL_RID" > "Finite_Set.AC_mult.f_e.ident" - "REAL_MUL_LZERO" > "Ring_and_Field.mult_zero_left" - "REAL_MUL_LNEG" > "Ring_and_Field.mult_minus_left" + "REAL_MUL_LZERO" > "Rings.mult_zero_left" + "REAL_MUL_LNEG" > "Rings.mult_minus_left" "REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV" "REAL_MUL_LID" > "Finite_Set.AC_mult.f_e.left_ident" "REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC" "REAL_MUL" > "RealDef.real_of_nat_mult" "REAL_MIDDLE2" > "HOL4Real.real.REAL_MIDDLE2" "REAL_MIDDLE1" > "HOL4Real.real.REAL_MIDDLE1" - "REAL_MEAN" > "Ring_and_Field.dense" + "REAL_MEAN" > "Rings.dense" "REAL_LT_TRANS" > "Set.basic_trans_rules_21" "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL" - "REAL_LT_SUB_RADD" > "OrderedGroup.compare_rls_6" - "REAL_LT_SUB_LADD" > "OrderedGroup.compare_rls_7" - "REAL_LT_RMUL_IMP" > "Ring_and_Field.mult_strict_right_mono" + "REAL_LT_SUB_RADD" > "Groups.compare_rls_6" + "REAL_LT_SUB_LADD" > "Groups.compare_rls_7" + "REAL_LT_RMUL_IMP" > "Rings.mult_strict_right_mono" "REAL_LT_RMUL_0" > "HOL4Real.real.REAL_LT_RMUL_0" "REAL_LT_RMUL" > "RealDef.real_mult_less_iff1" "REAL_LT_REFL" > "Orderings.order_less_irrefl" - "REAL_LT_RDIV_EQ" > "Ring_and_Field.pos_less_divide_eq" + "REAL_LT_RDIV_EQ" > "Rings.pos_less_divide_eq" "REAL_LT_RDIV_0" > "HOL4Real.real.REAL_LT_RDIV_0" "REAL_LT_RDIV" > "HOL4Real.real.REAL_LT_RDIV" - "REAL_LT_RADD" > "OrderedGroup.add_less_cancel_right" + "REAL_LT_RADD" > "Groups.add_less_cancel_right" "REAL_LT_NZ" > "HOL4Real.real.REAL_LT_NZ" "REAL_LT_NEGTOTAL" > "HOL4Real.real.REAL_LT_NEGTOTAL" - "REAL_LT_NEG" > "OrderedGroup.neg_less_iff_less" + "REAL_LT_NEG" > "Groups.neg_less_iff_less" "REAL_LT_MULTIPLE" > "HOL4Real.real.REAL_LT_MULTIPLE" - "REAL_LT_MUL2" > "Ring_and_Field.mult_strict_mono'" - "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos" - "REAL_LT_LMUL_IMP" > "Ring_and_Field.linordered_comm_semiring_strict_class.mult_strict_mono" + "REAL_LT_MUL2" > "Rings.mult_strict_mono'" + "REAL_LT_MUL" > "Rings.mult_pos_pos" + "REAL_LT_LMUL_IMP" > "Rings.linordered_comm_semiring_strict_class.mult_strict_mono" "REAL_LT_LMUL_0" > "HOL4Real.real.REAL_LT_LMUL_0" "REAL_LT_LMUL" > "HOL4Real.real.REAL_LT_LMUL" "REAL_LT_LE" > "Orderings.order_class.order_less_le" - "REAL_LT_LDIV_EQ" > "Ring_and_Field.pos_divide_less_eq" - "REAL_LT_LADD" > "OrderedGroup.add_less_cancel_left" - "REAL_LT_INV_EQ" > "Ring_and_Field.inverse_positive_iff_positive" - "REAL_LT_INV" > "Ring_and_Field.less_imp_inverse_less" + "REAL_LT_LDIV_EQ" > "Rings.pos_divide_less_eq" + "REAL_LT_LADD" > "Groups.add_less_cancel_left" + "REAL_LT_INV_EQ" > "Rings.inverse_positive_iff_positive" + "REAL_LT_INV" > "Rings.less_imp_inverse_less" "REAL_LT_IMP_NE" > "Orderings.less_imp_neq" "REAL_LT_IMP_LE" > "Orderings.order_less_imp_le" - "REAL_LT_IADD" > "OrderedGroup.add_strict_left_mono" + "REAL_LT_IADD" > "Groups.add_strict_left_mono" "REAL_LT_HALF2" > "HOL4Real.real.REAL_LT_HALF2" "REAL_LT_HALF1" > "NatSimprocs.half_gt_zero_iff" "REAL_LT_GT" > "Orderings.order_less_not_sym" "REAL_LT_FRACTION_0" > "HOL4Real.real.REAL_LT_FRACTION_0" "REAL_LT_FRACTION" > "HOL4Real.real.REAL_LT_FRACTION" - "REAL_LT_DIV" > "Ring_and_Field.divide_pos_pos" + "REAL_LT_DIV" > "Rings.divide_pos_pos" "REAL_LT_ANTISYM" > "HOL4Real.real.REAL_LT_ANTISYM" - "REAL_LT_ADD_SUB" > "OrderedGroup.compare_rls_7" + "REAL_LT_ADD_SUB" > "Groups.compare_rls_7" "REAL_LT_ADDR" > "HOL4Real.real.REAL_LT_ADDR" "REAL_LT_ADDNEG2" > "HOL4Real.real.REAL_LT_ADDNEG2" "REAL_LT_ADDNEG" > "HOL4Real.real.REAL_LT_ADDNEG" "REAL_LT_ADDL" > "HOL4Real.real.REAL_LT_ADDL" - "REAL_LT_ADD2" > "OrderedGroup.add_strict_mono" + "REAL_LT_ADD2" > "Groups.add_strict_mono" "REAL_LT_ADD1" > "HOL4Real.real.REAL_LT_ADD1" - "REAL_LT_ADD" > "OrderedGroup.add_pos_pos" + "REAL_LT_ADD" > "Groups.add_pos_pos" "REAL_LT_1" > "HOL4Real.real.REAL_LT_1" - "REAL_LT_01" > "Ring_and_Field.zero_less_one" + "REAL_LT_01" > "Rings.zero_less_one" "REAL_LTE_TRANS" > "Set.basic_trans_rules_24" "REAL_LTE_TOTAL" > "HOL4Real.real.REAL_LTE_TOTAL" "REAL_LTE_ANTSYM" > "HOL4Real.real.REAL_LTE_ANTSYM" - "REAL_LTE_ADD2" > "OrderedGroup.add_less_le_mono" - "REAL_LTE_ADD" > "OrderedGroup.add_pos_nonneg" + "REAL_LTE_ADD2" > "Groups.add_less_le_mono" + "REAL_LTE_ADD" > "Groups.add_pos_nonneg" "REAL_LT1_POW2" > "HOL4Real.real.REAL_LT1_POW2" "REAL_LT" > "RealDef.real_of_nat_less_iff" "REAL_LNEG_UNIQ" > "HOL4Real.real.REAL_LNEG_UNIQ" "REAL_LINV_UNIQ" > "HOL4Real.real.REAL_LINV_UNIQ" "REAL_LE_TRANS" > "Set.basic_trans_rules_25" "REAL_LE_TOTAL" > "Orderings.linorder_class.linorder_linear" - "REAL_LE_SUB_RADD" > "OrderedGroup.compare_rls_8" - "REAL_LE_SUB_LADD" > "OrderedGroup.compare_rls_9" - "REAL_LE_SQUARE" > "Ring_and_Field.zero_le_square" - "REAL_LE_RNEG" > "OrderedGroup.le_eq_neg" - "REAL_LE_RMUL_IMP" > "Ring_and_Field.mult_right_mono" + "REAL_LE_SUB_RADD" > "Groups.compare_rls_8" + "REAL_LE_SUB_LADD" > "Groups.compare_rls_9" + "REAL_LE_SQUARE" > "Rings.zero_le_square" + "REAL_LE_RNEG" > "Groups.le_eq_neg" + "REAL_LE_RMUL_IMP" > "Rings.mult_right_mono" "REAL_LE_RMUL" > "RealDef.real_mult_le_cancel_iff1" "REAL_LE_REFL" > "Finite_Set.max.f_below.below_refl" - "REAL_LE_RDIV_EQ" > "Ring_and_Field.pos_le_divide_eq" - "REAL_LE_RDIV" > "Ring_and_Field.mult_imp_le_div_pos" - "REAL_LE_RADD" > "OrderedGroup.add_le_cancel_right" + "REAL_LE_RDIV_EQ" > "Rings.pos_le_divide_eq" + "REAL_LE_RDIV" > "Rings.mult_imp_le_div_pos" + "REAL_LE_RADD" > "Groups.add_le_cancel_right" "REAL_LE_POW2" > "Nat_Numeral.zero_compare_simps_12" "REAL_LE_NEGTOTAL" > "HOL4Real.real.REAL_LE_NEGTOTAL" - "REAL_LE_NEGR" > "OrderedGroup.le_minus_self_iff" - "REAL_LE_NEGL" > "OrderedGroup.minus_le_self_iff" - "REAL_LE_NEG2" > "OrderedGroup.neg_le_iff_le" - "REAL_LE_NEG" > "OrderedGroup.neg_le_iff_le" + "REAL_LE_NEGR" > "Groups.le_minus_self_iff" + "REAL_LE_NEGL" > "Groups.minus_le_self_iff" + "REAL_LE_NEG2" > "Groups.neg_le_iff_le" + "REAL_LE_NEG" > "Groups.neg_le_iff_le" "REAL_LE_MUL2" > "HOL4Real.real.REAL_LE_MUL2" - "REAL_LE_MUL" > "Ring_and_Field.mult_nonneg_nonneg" + "REAL_LE_MUL" > "Rings.mult_nonneg_nonneg" "REAL_LE_LT" > "Orderings.order_le_less" "REAL_LE_LNEG" > "RealDef.real_0_le_add_iff" - "REAL_LE_LMUL_IMP" > "Ring_and_Field.mult_mono" + "REAL_LE_LMUL_IMP" > "Rings.mult_mono" "REAL_LE_LMUL" > "RealDef.real_mult_le_cancel_iff2" - "REAL_LE_LDIV_EQ" > "Ring_and_Field.pos_divide_le_eq" - "REAL_LE_LDIV" > "Ring_and_Field.mult_imp_div_pos_le" - "REAL_LE_LADD_IMP" > "OrderedGroup.add_left_mono" - "REAL_LE_LADD" > "OrderedGroup.add_le_cancel_left" - "REAL_LE_INV_EQ" > "Ring_and_Field.inverse_nonnegative_iff_nonnegative" + "REAL_LE_LDIV_EQ" > "Rings.pos_divide_le_eq" + "REAL_LE_LDIV" > "Rings.mult_imp_div_pos_le" + "REAL_LE_LADD_IMP" > "Groups.add_left_mono" + "REAL_LE_LADD" > "Groups.add_le_cancel_left" + "REAL_LE_INV_EQ" > "Rings.inverse_nonnegative_iff_nonnegative" "REAL_LE_INV" > "HOL4Real.real.REAL_LE_INV" - "REAL_LE_DOUBLE" > "OrderedGroup.zero_le_double_add_iff_zero_le_single_add" + "REAL_LE_DOUBLE" > "Groups.zero_le_double_add_iff_zero_le_single_add" "REAL_LE_DIV" > "HOL4Real.real.REAL_LE_DIV" "REAL_LE_ANTISYM" > "Orderings.order_eq_iff" "REAL_LE_ADDR" > "HOL4Real.real.REAL_LE_ADDR" "REAL_LE_ADDL" > "HOL4Real.real.REAL_LE_ADDL" - "REAL_LE_ADD2" > "OrderedGroup.add_mono" - "REAL_LE_ADD" > "OrderedGroup.add_nonneg_nonneg" - "REAL_LE_01" > "Ring_and_Field.zero_le_one" + "REAL_LE_ADD2" > "Groups.add_mono" + "REAL_LE_ADD" > "Groups.add_nonneg_nonneg" + "REAL_LE_01" > "Rings.zero_le_one" "REAL_LET_TRANS" > "Set.basic_trans_rules_23" "REAL_LET_TOTAL" > "Orderings.linorder_le_less_linear" "REAL_LET_ANTISYM" > "HOL4Real.real.REAL_LET_ANTISYM" - "REAL_LET_ADD2" > "OrderedGroup.add_le_less_mono" - "REAL_LET_ADD" > "OrderedGroup.add_nonneg_pos" + "REAL_LET_ADD2" > "Groups.add_le_less_mono" + "REAL_LET_ADD" > "Groups.add_nonneg_pos" "REAL_LE1_POW2" > "HOL4Real.real.REAL_LE1_POW2" "REAL_LE" > "RealDef.real_of_nat_le_iff" - "REAL_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2" - "REAL_INV_POS" > "Ring_and_Field.positive_imp_inverse_positive" - "REAL_INV_NZ" > "Ring_and_Field.nonzero_imp_inverse_nonzero" + "REAL_LDISTRIB" > "Rings.ring_eq_simps_2" + "REAL_INV_POS" > "Rings.positive_imp_inverse_positive" + "REAL_INV_NZ" > "Rings.nonzero_imp_inverse_nonzero" "REAL_INV_MUL" > "HOL4Real.real.REAL_INV_MUL" "REAL_INV_LT1" > "RealDef.real_inverse_gt_one" - "REAL_INV_INV" > "Ring_and_Field.inverse_inverse_eq" - "REAL_INV_EQ_0" > "Ring_and_Field.inverse_nonzero_iff_nonzero" - "REAL_INV_1OVER" > "Ring_and_Field.inverse_eq_divide" - "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero" - "REAL_INVINV" > "Ring_and_Field.nonzero_inverse_inverse_eq" - "REAL_INV1" > "Ring_and_Field.inverse_1" + "REAL_INV_INV" > "Rings.inverse_inverse_eq" + "REAL_INV_EQ_0" > "Rings.inverse_nonzero_iff_nonzero" + "REAL_INV_1OVER" > "Rings.inverse_eq_divide" + "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero" + "REAL_INVINV" > "Rings.nonzero_inverse_inverse_eq" + "REAL_INV1" > "Rings.inverse_1" "REAL_INJ" > "RealDef.real_of_nat_inject" "REAL_HALF_DOUBLE" > "RComplete.real_sum_of_halves" "REAL_FACT_NZ" > "HOL4Real.real.REAL_FACT_NZ" - "REAL_EQ_SUB_RADD" > "Ring_and_Field.ring_eq_simps_15" - "REAL_EQ_SUB_LADD" > "Ring_and_Field.ring_eq_simps_16" - "REAL_EQ_RMUL_IMP" > "Ring_and_Field.field_mult_cancel_right_lemma" - "REAL_EQ_RMUL" > "Ring_and_Field.field_mult_cancel_right" + "REAL_EQ_SUB_RADD" > "Rings.ring_eq_simps_15" + "REAL_EQ_SUB_LADD" > "Rings.ring_eq_simps_16" + "REAL_EQ_RMUL_IMP" > "Rings.field_mult_cancel_right_lemma" + "REAL_EQ_RMUL" > "Rings.field_mult_cancel_right" "REAL_EQ_RDIV_EQ" > "HOL4Real.real.REAL_EQ_RDIV_EQ" - "REAL_EQ_RADD" > "OrderedGroup.add_right_cancel" - "REAL_EQ_NEG" > "OrderedGroup.neg_equal_iff_equal" - "REAL_EQ_MUL_LCANCEL" > "Ring_and_Field.field_mult_cancel_left" + "REAL_EQ_RADD" > "Groups.add_right_cancel" + "REAL_EQ_NEG" > "Groups.neg_equal_iff_equal" + "REAL_EQ_MUL_LCANCEL" > "Rings.field_mult_cancel_left" "REAL_EQ_LMUL_IMP" > "HOL4Real.real.REAL_EQ_LMUL_IMP" "REAL_EQ_LMUL2" > "RealDef.real_mult_left_cancel" - "REAL_EQ_LMUL" > "Ring_and_Field.field_mult_cancel_left" + "REAL_EQ_LMUL" > "Rings.field_mult_cancel_left" "REAL_EQ_LDIV_EQ" > "HOL4Real.real.REAL_EQ_LDIV_EQ" - "REAL_EQ_LADD" > "OrderedGroup.add_left_cancel" + "REAL_EQ_LADD" > "Groups.add_left_cancel" "REAL_EQ_IMP_LE" > "Orderings.order_eq_refl" - "REAL_ENTIRE" > "Ring_and_Field.field_mult_eq_0_iff" + "REAL_ENTIRE" > "Rings.field_mult_eq_0_iff" "REAL_DOWN2" > "RealDef.real_lbound_gt_zero" "REAL_DOWN" > "HOL4Real.real.REAL_DOWN" "REAL_DOUBLE" > "Int.mult_2" "REAL_DIV_RMUL" > "HOL4Real.real.REAL_DIV_RMUL" - "REAL_DIV_REFL" > "Ring_and_Field.divide_self" + "REAL_DIV_REFL" > "Rings.divide_self" "REAL_DIV_MUL2" > "HOL4Real.real.REAL_DIV_MUL2" - "REAL_DIV_LZERO" > "Ring_and_Field.divide_zero_left" + "REAL_DIV_LZERO" > "Rings.divide_zero_left" "REAL_DIV_LMUL" > "HOL4Real.real.REAL_DIV_LMUL" "REAL_DIFFSQ" > "HOL4Real.real.REAL_DIFFSQ" "REAL_ARCH_LEAST" > "HOL4Real.real.REAL_ARCH_LEAST" @@ -286,20 +286,20 @@ "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2" "REAL_ADD_SUB2" > "HOL4Real.real.REAL_ADD_SUB2" "REAL_ADD_SUB" > "HOL4Real.real.REAL_ADD_SUB" - "REAL_ADD_RINV" > "OrderedGroup.right_minus" + "REAL_ADD_RINV" > "Groups.right_minus" "REAL_ADD_RID_UNIQ" > "HOL4Real.real.REAL_ADD_RID_UNIQ" "REAL_ADD_RID" > "Finite_Set.AC_add.f_e.ident" - "REAL_ADD_RDISTRIB" > "Ring_and_Field.ring_eq_simps_1" + "REAL_ADD_RDISTRIB" > "Rings.ring_eq_simps_1" "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV" "REAL_ADD_LID_UNIQ" > "HOL4Real.real.REAL_ADD_LID_UNIQ" "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident" - "REAL_ADD_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2" + "REAL_ADD_LDISTRIB" > "Rings.ring_eq_simps_2" "REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC" "REAL_ADD2_SUB2" > "HOL4Real.real.REAL_ADD2_SUB2" "REAL_ADD" > "RealDef.real_of_nat_add" - "REAL_ABS_TRIANGLE" > "OrderedGroup.abs_triangle_ineq" - "REAL_ABS_POS" > "OrderedGroup.abs_ge_zero" - "REAL_ABS_MUL" > "Ring_and_Field.abs_mult" + "REAL_ABS_TRIANGLE" > "Groups.abs_triangle_ineq" + "REAL_ABS_POS" > "Groups.abs_ge_zero" + "REAL_ABS_MUL" > "Rings.abs_mult" "REAL_ABS_0" > "Int.bin_arith_simps_28" "REAL_10" > "HOL4Compat.REAL_10" "REAL_1" > "HOL4Real.real.REAL_1" @@ -326,25 +326,25 @@ "POW_2" > "Nat_Numeral.power2_eq_square" "POW_1" > "Power.power_one_right" "POW_0" > "Power.power_0_Suc" - "ABS_ZERO" > "OrderedGroup.abs_eq_0" - "ABS_TRIANGLE" > "OrderedGroup.abs_triangle_ineq" + "ABS_ZERO" > "Groups.abs_eq_0" + "ABS_TRIANGLE" > "Groups.abs_triangle_ineq" "ABS_SUM" > "HOL4Real.real.ABS_SUM" - "ABS_SUB_ABS" > "OrderedGroup.abs_triangle_ineq3" - "ABS_SUB" > "OrderedGroup.abs_minus_commute" + "ABS_SUB_ABS" > "Groups.abs_triangle_ineq3" + "ABS_SUB" > "Groups.abs_minus_commute" "ABS_STILLNZ" > "HOL4Real.real.ABS_STILLNZ" "ABS_SIGN2" > "HOL4Real.real.ABS_SIGN2" "ABS_SIGN" > "HOL4Real.real.ABS_SIGN" "ABS_REFL" > "HOL4Real.real.ABS_REFL" "ABS_POW2" > "Nat_Numeral.abs_power2" - "ABS_POS" > "OrderedGroup.abs_ge_zero" - "ABS_NZ" > "OrderedGroup.zero_less_abs_iff" - "ABS_NEG" > "OrderedGroup.abs_minus_cancel" + "ABS_POS" > "Groups.abs_ge_zero" + "ABS_NZ" > "Groups.zero_less_abs_iff" + "ABS_NEG" > "Groups.abs_minus_cancel" "ABS_N" > "RealDef.abs_real_of_nat_cancel" - "ABS_MUL" > "Ring_and_Field.abs_mult" + "ABS_MUL" > "Rings.abs_mult" "ABS_LT_MUL2" > "HOL4Real.real.ABS_LT_MUL2" - "ABS_LE" > "OrderedGroup.abs_ge_self" - "ABS_INV" > "Ring_and_Field.nonzero_abs_inverse" - "ABS_DIV" > "Ring_and_Field.nonzero_abs_divide" + "ABS_LE" > "Groups.abs_ge_self" + "ABS_INV" > "Rings.nonzero_abs_inverse" + "ABS_DIV" > "Rings.nonzero_abs_divide" "ABS_CIRCLE" > "HOL4Real.real.ABS_CIRCLE" "ABS_CASES" > "HOL4Real.real.ABS_CASES" "ABS_BOUNDS" > "RealDef.abs_le_interval_iff" @@ -352,7 +352,7 @@ "ABS_BETWEEN2" > "HOL4Real.real.ABS_BETWEEN2" "ABS_BETWEEN1" > "HOL4Real.real.ABS_BETWEEN1" "ABS_BETWEEN" > "HOL4Real.real.ABS_BETWEEN" - "ABS_ABS" > "OrderedGroup.abs_idempotent" + "ABS_ABS" > "Groups.abs_idempotent" "ABS_1" > "Int.bin_arith_simps_29" "ABS_0" > "Int.bin_arith_simps_28" diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Import/HOL/realax.imp --- a/src/HOL/Import/HOL/realax.imp Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Import/HOL/realax.imp Mon Feb 08 17:12:38 2010 +0100 @@ -98,10 +98,10 @@ "REAL_LT_TRANS" > "Set.basic_trans_rules_21" "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL" "REAL_LT_REFL" > "Orderings.order_less_irrefl" - "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos" - "REAL_LT_IADD" > "OrderedGroup.add_strict_left_mono" - "REAL_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2" - "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero" + "REAL_LT_MUL" > "Rings.mult_pos_pos" + "REAL_LT_IADD" > "Groups.add_strict_left_mono" + "REAL_LDISTRIB" > "Rings.ring_eq_simps_2" + "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero" "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2" "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV" "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident" diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Int.thy --- a/src/HOL/Int.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Int.thy Mon Feb 08 17:12:38 2010 +0100 @@ -307,7 +307,7 @@ by (cases z, simp add: algebra_simps of_int minus) lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" -by (simp add: OrderedGroup.diff_minus diff_minus) +by (simp add: diff_minus Groups.diff_minus) lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" apply (cases w, cases z) @@ -519,7 +519,7 @@ text{*This version is proved for all ordered rings, not just integers! It is proved here because attribute @{text arith_split} is not available - in theory @{text Ring_and_Field}. + in theory @{text Rings}. But is it really better than just rewriting with @{text abs_if}?*} lemma abs_split [arith_split,noatp]: "P(abs(a::'a::linordered_idom)) = ((0 \ a --> P a) & (a < 0 --> P(-a)))" @@ -2317,9 +2317,9 @@ lemmas zadd_assoc = add_assoc [of "z1::int" "z2" "z3", standard] lemmas zadd_left_commute = add_left_commute [of "x::int" "y" "z", standard] lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute -lemmas zmult_ac = OrderedGroup.mult_ac -lemmas zadd_0 = OrderedGroup.add_0_left [of "z::int", standard] -lemmas zadd_0_right = OrderedGroup.add_0_left [of "z::int", standard] +lemmas zmult_ac = mult_ac +lemmas zadd_0 = add_0_left [of "z::int", standard] +lemmas zadd_0_right = add_0_right [of "z::int", standard] lemmas zadd_zminus_inverse2 = left_minus [of "z::int", standard] lemmas zmult_zminus = mult_minus_left [of "z::int" "w", standard] lemmas zmult_commute = mult_commute [of "z::int" "w", standard] diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/IsaMakefile Mon Feb 08 17:12:38 2010 +0100 @@ -145,15 +145,16 @@ Complete_Lattice.thy \ Datatype.thy \ Extraction.thy \ + Fields.thy \ Finite_Set.thy \ Fun.thy \ FunDef.thy \ + Groups.thy \ Inductive.thy \ Lattices.thy \ Nat.thy \ Nitpick.thy \ Option.thy \ - OrderedGroup.thy \ Orderings.thy \ Plain.thy \ Power.thy \ @@ -162,7 +163,7 @@ Record.thy \ Refute.thy \ Relation.thy \ - Ring_and_Field.thy \ + Rings.thy \ SAT.thy \ Set.thy \ Sum_Type.thy \ diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Library/Poly_Deriv.thy --- a/src/HOL/Library/Poly_Deriv.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Library/Poly_Deriv.thy Mon Feb 08 17:12:38 2010 +0100 @@ -139,7 +139,7 @@ lemma dvd_add_cancel1: fixes a b c :: "'a::comm_ring_1" shows "a dvd b + c \ a dvd b \ a dvd c" - by (drule (1) Ring_and_Field.dvd_diff, simp) + by (drule (1) Rings.dvd_diff, simp) lemma lemma_order_pderiv [rule_format]: "\p q a. 0 < n & diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Metis_Examples/BigO.thy --- a/src/HOL/Metis_Examples/BigO.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Metis_Examples/BigO.thy Mon Feb 08 17:12:38 2010 +0100 @@ -368,7 +368,7 @@ f : O(g)" apply (auto simp add: bigo_def) (*Version 1: one-shot proof*) - apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult) + apply (metis abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) abs_mult) done lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> @@ -595,8 +595,8 @@ using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]] prefer 2 apply (metis mult_assoc mult_left_commute - OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute - Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos) + abs_of_pos mult_left_commute + abs_mult mult_pos_pos) apply (erule ssubst) apply (subst abs_mult) (*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has @@ -613,32 +613,32 @@ \ (c\'b\linordered_idom) * \(f\'a \ 'b\linordered_idom) x\ * ((ca\'b\linordered_idom) * \(g\'a \ 'b\linordered_idom) x\)" have 4: "\c\'b\linordered_idom\ = c" - by (metis OrderedGroup.abs_of_pos 0) + by (metis abs_of_pos 0) have 5: "\X1\'b\linordered_idom. (c\'b\linordered_idom) * \X1\ = \c * X1\" - by (metis Ring_and_Field.abs_mult 4) + by (metis abs_mult 4) have 6: "(0\'b\linordered_idom) = (1\'b\linordered_idom) \ (0\'b\linordered_idom) < (1\'b\linordered_idom)" - by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_linordered_idom) + by (metis abs_not_less_zero abs_one linorder_neqE_linordered_idom) have 7: "(0\'b\linordered_idom) < (1\'b\linordered_idom)" - by (metis 6 Ring_and_Field.one_neq_zero) + by (metis 6 one_neq_zero) have 8: "\1\'b\linordered_idom\ = (1\'b\linordered_idom)" - by (metis OrderedGroup.abs_of_pos 7) + by (metis abs_of_pos 7) have 9: "\X1\'b\linordered_idom. (0\'b\linordered_idom) \ (c\'b\linordered_idom) * \X1\" - by (metis OrderedGroup.abs_ge_zero 5) + by (metis abs_ge_zero 5) have 10: "\X1\'b\linordered_idom. X1 * (1\'b\linordered_idom) = X1" - by (metis Ring_and_Field.mult_cancel_right2 mult_commute) + by (metis mult_cancel_right2 mult_commute) have 11: "\X1\'b\linordered_idom. \\X1\\ = \X1\ * \1\'b\linordered_idom\" - by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10) + by (metis abs_mult abs_idempotent 10) have 12: "\X1\'b\linordered_idom. \\X1\\ = \X1\" by (metis 11 8 10) have 13: "\X1\'b\linordered_idom. (0\'b\linordered_idom) \ \X1\" - by (metis OrderedGroup.abs_ge_zero 12) + by (metis abs_ge_zero 12) have 14: "\ (0\'b\linordered_idom) \ (c\'b\linordered_idom) * \(f\'a \ 'b\linordered_idom) (x\'a)\ \ \ (0\'b\linordered_idom) \ \(b\'a \ 'b\linordered_idom) x\ \ \ \b x\ \ (ca\'b\linordered_idom) * \(g\'a \ 'b\linordered_idom) x\ \ \ \(a\'a \ 'b\linordered_idom) x\ \ c * \f x\" - by (metis 3 Ring_and_Field.mult_mono) + by (metis 3 mult_mono) have 15: "\ (0\'b\linordered_idom) \ \(b\'a \ 'b\linordered_idom) (x\'a)\ \ \ \b x\ \ (ca\'b\linordered_idom) * \(g\'a \ 'b\linordered_idom) x\ \ \ \(a\'a \ 'b\linordered_idom) x\ @@ -1078,7 +1078,7 @@ apply (rule_tac x = c in exI) apply safe apply (case_tac "x = 0") -apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le) +apply (metis abs_ge_zero abs_zero order_less_le split_mult_pos_le) apply (subgoal_tac "x = Suc (x - 1)") apply metis apply simp diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/NSA/HyperDef.thy --- a/src/HOL/NSA/HyperDef.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/NSA/HyperDef.thy Mon Feb 08 17:12:38 2010 +0100 @@ -394,7 +394,7 @@ by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two) (*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract - result proved in Ring_and_Field*) + result proved in Rings or Fields*) lemma hrabs_hrealpow_two [simp]: "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)" by (simp add: abs_mult) @@ -516,7 +516,7 @@ text{*The precondition could be weakened to @{term "0\x"}*} lemma hypreal_mult_less_mono: "[| u u*x < v* y" - by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) + by (simp add: mult_strict_mono order_less_imp_le) lemma hyperpow_two_gt_one: "\r::'a::{linordered_semidom} star. 1 < r \ 1 < r pow (1 + 1)" diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/NSA/NSComplex.thy --- a/src/HOL/NSA/NSComplex.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/NSA/NSComplex.thy Mon Feb 08 17:12:38 2010 +0100 @@ -1,5 +1,4 @@ (* Title: NSComplex.thy - ID: $Id$ Author: Jacques D. Fleuriot Copyright: 2001 University of Edinburgh Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 @@ -161,7 +160,7 @@ lemma hcomplex_add_minus_eq_minus: "x + y = (0::hcomplex) ==> x = -y" -apply (drule OrderedGroup.minus_unique) +apply (drule minus_unique) apply (simp add: minus_equation_iff [of x y]) done @@ -196,7 +195,7 @@ lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)" (* TODO: delete *) -by (rule OrderedGroup.diff_eq_eq) +by (rule diff_eq_eq) subsection{*Embedding Properties for @{term hcomplex_of_hypreal} Map*} diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/NSA/StarDef.thy Mon Feb 08 17:12:38 2010 +0100 @@ -530,7 +530,7 @@ end -instance star :: (Ring_and_Field.dvd) Ring_and_Field.dvd .. +instance star :: (Rings.dvd) Rings.dvd .. instantiation star :: (Divides.div) Divides.div begin diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Mon Feb 08 17:12:32 2010 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1159 +0,0 @@ -(* Title: HOL/OrderedGroup.thy - Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad -*) - -header {* Ordered Groups *} - -theory OrderedGroup -imports Lattices -uses "~~/src/Provers/Arith/abel_cancel.ML" -begin - -text {* - The theory of partially ordered groups is taken from the books: - \begin{itemize} - \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 - \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963 - \end{itemize} - Most of the used notions can also be looked up in - \begin{itemize} - \item \url{http://www.mathworld.com} by Eric Weisstein et. al. - \item \emph{Algebra I} by van der Waerden, Springer. - \end{itemize} -*} - -ML {* -structure Algebra_Simps = Named_Thms( - val name = "algebra_simps" - val description = "algebra simplification rules" -) -*} - -setup Algebra_Simps.setup - -text{* The rewrites accumulated in @{text algebra_simps} deal with the -classical algebraic structures of groups, rings and family. They simplify -terms by multiplying everything out (in case of a ring) and bringing sums and -products into a canonical form (by ordered rewriting). As a result it decides -group and ring equalities but also helps with inequalities. - -Of course it also works for fields, but it knows nothing about multiplicative -inverses or division. This is catered for by @{text field_simps}. *} - -subsection {* Semigroups and Monoids *} - -class semigroup_add = plus + - assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)" - -sublocale semigroup_add < plus!: semigroup plus proof -qed (fact add_assoc) - -class ab_semigroup_add = semigroup_add + - assumes add_commute [algebra_simps]: "a + b = b + a" - -sublocale ab_semigroup_add < plus!: abel_semigroup plus proof -qed (fact add_commute) - -context ab_semigroup_add -begin - -lemmas add_left_commute [algebra_simps] = plus.left_commute - -theorems add_ac = add_assoc add_commute add_left_commute - -end - -theorems add_ac = add_assoc add_commute add_left_commute - -class semigroup_mult = times + - assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)" - -sublocale semigroup_mult < times!: semigroup times proof -qed (fact mult_assoc) - -class ab_semigroup_mult = semigroup_mult + - assumes mult_commute [algebra_simps]: "a * b = b * a" - -sublocale ab_semigroup_mult < times!: abel_semigroup times proof -qed (fact mult_commute) - -context ab_semigroup_mult -begin - -lemmas mult_left_commute [algebra_simps] = times.left_commute - -theorems mult_ac = mult_assoc mult_commute mult_left_commute - -end - -theorems mult_ac = mult_assoc mult_commute mult_left_commute - -class ab_semigroup_idem_mult = ab_semigroup_mult + - assumes mult_idem: "x * x = x" - -sublocale ab_semigroup_idem_mult < times!: semilattice times proof -qed (fact mult_idem) - -context ab_semigroup_idem_mult -begin - -lemmas mult_left_idem = times.left_idem - -end - -class monoid_add = zero + semigroup_add + - assumes add_0_left [simp]: "0 + a = a" - and add_0_right [simp]: "a + 0 = a" - -lemma zero_reorient: "0 = x \ x = 0" -by (rule eq_commute) - -class comm_monoid_add = zero + ab_semigroup_add + - assumes add_0: "0 + a = a" -begin - -subclass monoid_add - proof qed (insert add_0, simp_all add: add_commute) - -end - -class monoid_mult = one + semigroup_mult + - assumes mult_1_left [simp]: "1 * a = a" - assumes mult_1_right [simp]: "a * 1 = a" - -lemma one_reorient: "1 = x \ x = 1" -by (rule eq_commute) - -class comm_monoid_mult = one + ab_semigroup_mult + - assumes mult_1: "1 * a = a" -begin - -subclass monoid_mult - proof qed (insert mult_1, simp_all add: mult_commute) - -end - -class cancel_semigroup_add = semigroup_add + - assumes add_left_imp_eq: "a + b = a + c \ b = c" - assumes add_right_imp_eq: "b + a = c + a \ b = c" -begin - -lemma add_left_cancel [simp]: - "a + b = a + c \ b = c" -by (blast dest: add_left_imp_eq) - -lemma add_right_cancel [simp]: - "b + a = c + a \ b = c" -by (blast dest: add_right_imp_eq) - -end - -class cancel_ab_semigroup_add = ab_semigroup_add + - assumes add_imp_eq: "a + b = a + c \ b = c" -begin - -subclass cancel_semigroup_add -proof - fix a b c :: 'a - assume "a + b = a + c" - then show "b = c" by (rule add_imp_eq) -next - fix a b c :: 'a - assume "b + a = c + a" - then have "a + b = a + c" by (simp only: add_commute) - then show "b = c" by (rule add_imp_eq) -qed - -end - -class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add - - -subsection {* Groups *} - -class group_add = minus + uminus + monoid_add + - assumes left_minus [simp]: "- a + a = 0" - assumes diff_minus: "a - b = a + (- b)" -begin - -lemma minus_unique: - assumes "a + b = 0" shows "- a = b" -proof - - have "- a = - a + (a + b)" using assms by simp - also have "\ = b" by (simp add: add_assoc [symmetric]) - finally show ?thesis . -qed - -lemmas equals_zero_I = minus_unique (* legacy name *) - -lemma minus_zero [simp]: "- 0 = 0" -proof - - have "0 + 0 = 0" by (rule add_0_right) - thus "- 0 = 0" by (rule minus_unique) -qed - -lemma minus_minus [simp]: "- (- a) = a" -proof - - have "- a + a = 0" by (rule left_minus) - thus "- (- a) = a" by (rule minus_unique) -qed - -lemma right_minus [simp]: "a + - a = 0" -proof - - have "a + - a = - (- a) + - a" by simp - also have "\ = 0" by (rule left_minus) - finally show ?thesis . -qed - -lemma minus_add_cancel: "- a + (a + b) = b" -by (simp add: add_assoc [symmetric]) - -lemma add_minus_cancel: "a + (- a + b) = b" -by (simp add: add_assoc [symmetric]) - -lemma minus_add: "- (a + b) = - b + - a" -proof - - have "(a + b) + (- b + - a) = 0" - by (simp add: add_assoc add_minus_cancel) - thus "- (a + b) = - b + - a" - by (rule minus_unique) -qed - -lemma right_minus_eq: "a - b = 0 \ a = b" -proof - assume "a - b = 0" - have "a = (a - b) + b" by (simp add:diff_minus add_assoc) - also have "\ = b" using `a - b = 0` by simp - finally show "a = b" . -next - assume "a = b" thus "a - b = 0" by (simp add: diff_minus) -qed - -lemma diff_self [simp]: "a - a = 0" -by (simp add: diff_minus) - -lemma diff_0 [simp]: "0 - a = - a" -by (simp add: diff_minus) - -lemma diff_0_right [simp]: "a - 0 = a" -by (simp add: diff_minus) - -lemma diff_minus_eq_add [simp]: "a - - b = a + b" -by (simp add: diff_minus) - -lemma neg_equal_iff_equal [simp]: - "- a = - b \ a = b" -proof - assume "- a = - b" - hence "- (- a) = - (- b)" by simp - thus "a = b" by simp -next - assume "a = b" - thus "- a = - b" by simp -qed - -lemma neg_equal_0_iff_equal [simp]: - "- a = 0 \ a = 0" -by (subst neg_equal_iff_equal [symmetric], simp) - -lemma neg_0_equal_iff_equal [simp]: - "0 = - a \ 0 = a" -by (subst neg_equal_iff_equal [symmetric], simp) - -text{*The next two equations can make the simplifier loop!*} - -lemma equation_minus_iff: - "a = - b \ b = - a" -proof - - have "- (- a) = - b \ - a = b" by (rule neg_equal_iff_equal) - thus ?thesis by (simp add: eq_commute) -qed - -lemma minus_equation_iff: - "- a = b \ - b = a" -proof - - have "- a = - (- b) \ a = -b" by (rule neg_equal_iff_equal) - thus ?thesis by (simp add: eq_commute) -qed - -lemma diff_add_cancel: "a - b + b = a" -by (simp add: diff_minus add_assoc) - -lemma add_diff_cancel: "a + b - b = a" -by (simp add: diff_minus add_assoc) - -declare diff_minus[symmetric, algebra_simps] - -lemma eq_neg_iff_add_eq_0: "a = - b \ a + b = 0" -proof - assume "a = - b" then show "a + b = 0" by simp -next - assume "a + b = 0" - moreover have "a + (b + - b) = (a + b) + - b" - by (simp only: add_assoc) - ultimately show "a = - b" by simp -qed - -end - -class ab_group_add = minus + uminus + comm_monoid_add + - assumes ab_left_minus: "- a + a = 0" - assumes ab_diff_minus: "a - b = a + (- b)" -begin - -subclass group_add - proof qed (simp_all add: ab_left_minus ab_diff_minus) - -subclass cancel_comm_monoid_add -proof - fix a b c :: 'a - assume "a + b = a + c" - then have "- a + a + b = - a + a + c" - unfolding add_assoc by simp - then show "b = c" by simp -qed - -lemma uminus_add_conv_diff[algebra_simps]: - "- a + b = b - a" -by (simp add:diff_minus add_commute) - -lemma minus_add_distrib [simp]: - "- (a + b) = - a + - b" -by (rule minus_unique) (simp add: add_ac) - -lemma minus_diff_eq [simp]: - "- (a - b) = b - a" -by (simp add: diff_minus add_commute) - -lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c" -by (simp add: diff_minus add_ac) - -lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b" -by (simp add: diff_minus add_ac) - -lemma diff_eq_eq[algebra_simps]: "a - b = c \ a = c + b" -by (auto simp add: diff_minus add_assoc) - -lemma eq_diff_eq[algebra_simps]: "a = c - b \ a + b = c" -by (auto simp add: diff_minus add_assoc) - -lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)" -by (simp add: diff_minus add_ac) - -lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b" -by (simp add: diff_minus add_ac) - -lemma eq_iff_diff_eq_0: "a = b \ a - b = 0" -by (simp add: algebra_simps) - -lemma diff_eq_0_iff_eq [simp, noatp]: "a - b = 0 \ a = b" -by (simp add: algebra_simps) - -end - -subsection {* (Partially) Ordered Groups *} - -class ordered_ab_semigroup_add = order + ab_semigroup_add + - assumes add_left_mono: "a \ b \ c + a \ c + b" -begin - -lemma add_right_mono: - "a \ b \ a + c \ b + c" -by (simp add: add_commute [of _ c] add_left_mono) - -text {* non-strict, in both arguments *} -lemma add_mono: - "a \ b \ c \ d \ a + c \ b + d" - apply (erule add_right_mono [THEN order_trans]) - apply (simp add: add_commute add_left_mono) - done - -end - -class ordered_cancel_ab_semigroup_add = - ordered_ab_semigroup_add + cancel_ab_semigroup_add -begin - -lemma add_strict_left_mono: - "a < b \ c + a < c + b" -by (auto simp add: less_le add_left_mono) - -lemma add_strict_right_mono: - "a < b \ a + c < b + c" -by (simp add: add_commute [of _ c] add_strict_left_mono) - -text{*Strict monotonicity in both arguments*} -lemma add_strict_mono: - "a < b \ c < d \ a + c < b + d" -apply (erule add_strict_right_mono [THEN less_trans]) -apply (erule add_strict_left_mono) -done - -lemma add_less_le_mono: - "a < b \ c \ d \ a + c < b + d" -apply (erule add_strict_right_mono [THEN less_le_trans]) -apply (erule add_left_mono) -done - -lemma add_le_less_mono: - "a \ b \ c < d \ a + c < b + d" -apply (erule add_right_mono [THEN le_less_trans]) -apply (erule add_strict_left_mono) -done - -end - -class ordered_ab_semigroup_add_imp_le = - ordered_cancel_ab_semigroup_add + - assumes add_le_imp_le_left: "c + a \ c + b \ a \ b" -begin - -lemma add_less_imp_less_left: - assumes less: "c + a < c + b" shows "a < b" -proof - - from less have le: "c + a <= c + b" by (simp add: order_le_less) - have "a <= b" - apply (insert le) - apply (drule add_le_imp_le_left) - by (insert le, drule add_le_imp_le_left, assumption) - moreover have "a \ b" - proof (rule ccontr) - assume "~(a \ b)" - then have "a = b" by simp - then have "c + a = c + b" by simp - with less show "False"by simp - qed - ultimately show "a < b" by (simp add: order_le_less) -qed - -lemma add_less_imp_less_right: - "a + c < b + c \ a < b" -apply (rule add_less_imp_less_left [of c]) -apply (simp add: add_commute) -done - -lemma add_less_cancel_left [simp]: - "c + a < c + b \ a < b" -by (blast intro: add_less_imp_less_left add_strict_left_mono) - -lemma add_less_cancel_right [simp]: - "a + c < b + c \ a < b" -by (blast intro: add_less_imp_less_right add_strict_right_mono) - -lemma add_le_cancel_left [simp]: - "c + a \ c + b \ a \ b" -by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) - -lemma add_le_cancel_right [simp]: - "a + c \ b + c \ a \ b" -by (simp add: add_commute [of a c] add_commute [of b c]) - -lemma add_le_imp_le_right: - "a + c \ b + c \ a \ b" -by simp - -lemma max_add_distrib_left: - "max x y + z = max (x + z) (y + z)" - unfolding max_def by auto - -lemma min_add_distrib_left: - "min x y + z = min (x + z) (y + z)" - unfolding min_def by auto - -end - -subsection {* Support for reasoning about signs *} - -class ordered_comm_monoid_add = - ordered_cancel_ab_semigroup_add + comm_monoid_add -begin - -lemma add_pos_nonneg: - assumes "0 < a" and "0 \ b" shows "0 < a + b" -proof - - have "0 + 0 < a + b" - using assms by (rule add_less_le_mono) - then show ?thesis by simp -qed - -lemma add_pos_pos: - assumes "0 < a" and "0 < b" shows "0 < a + b" -by (rule add_pos_nonneg) (insert assms, auto) - -lemma add_nonneg_pos: - assumes "0 \ a" and "0 < b" shows "0 < a + b" -proof - - have "0 + 0 < a + b" - using assms by (rule add_le_less_mono) - then show ?thesis by simp -qed - -lemma add_nonneg_nonneg: - assumes "0 \ a" and "0 \ b" shows "0 \ a + b" -proof - - have "0 + 0 \ a + b" - using assms by (rule add_mono) - then show ?thesis by simp -qed - -lemma add_neg_nonpos: - assumes "a < 0" and "b \ 0" shows "a + b < 0" -proof - - have "a + b < 0 + 0" - using assms by (rule add_less_le_mono) - then show ?thesis by simp -qed - -lemma add_neg_neg: - assumes "a < 0" and "b < 0" shows "a + b < 0" -by (rule add_neg_nonpos) (insert assms, auto) - -lemma add_nonpos_neg: - assumes "a \ 0" and "b < 0" shows "a + b < 0" -proof - - have "a + b < 0 + 0" - using assms by (rule add_le_less_mono) - then show ?thesis by simp -qed - -lemma add_nonpos_nonpos: - assumes "a \ 0" and "b \ 0" shows "a + b \ 0" -proof - - have "a + b \ 0 + 0" - using assms by (rule add_mono) - then show ?thesis by simp -qed - -lemmas add_sign_intros = - add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg - add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos - -lemma add_nonneg_eq_0_iff: - assumes x: "0 \ x" and y: "0 \ y" - shows "x + y = 0 \ x = 0 \ y = 0" -proof (intro iffI conjI) - have "x = x + 0" by simp - also have "x + 0 \ x + y" using y by (rule add_left_mono) - also assume "x + y = 0" - also have "0 \ x" using x . - finally show "x = 0" . -next - have "y = 0 + y" by simp - also have "0 + y \ x + y" using x by (rule add_right_mono) - also assume "x + y = 0" - also have "0 \ y" using y . - finally show "y = 0" . -next - assume "x = 0 \ y = 0" - then show "x + y = 0" by simp -qed - -end - -class ordered_ab_group_add = - ab_group_add + ordered_ab_semigroup_add -begin - -subclass ordered_cancel_ab_semigroup_add .. - -subclass ordered_ab_semigroup_add_imp_le -proof - fix a b c :: 'a - assume "c + a \ c + b" - hence "(-c) + (c + a) \ (-c) + (c + b)" by (rule add_left_mono) - hence "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add_assoc) - thus "a \ b" by simp -qed - -subclass ordered_comm_monoid_add .. - -lemma max_diff_distrib_left: - shows "max x y - z = max (x - z) (y - z)" -by (simp add: diff_minus, rule max_add_distrib_left) - -lemma min_diff_distrib_left: - shows "min x y - z = min (x - z) (y - z)" -by (simp add: diff_minus, rule min_add_distrib_left) - -lemma le_imp_neg_le: - assumes "a \ b" shows "-b \ -a" -proof - - have "-a+a \ -a+b" using `a \ b` by (rule add_left_mono) - hence "0 \ -a+b" by simp - hence "0 + (-b) \ (-a + b) + (-b)" by (rule add_right_mono) - thus ?thesis by (simp add: add_assoc) -qed - -lemma neg_le_iff_le [simp]: "- b \ - a \ a \ b" -proof - assume "- b \ - a" - hence "- (- a) \ - (- b)" by (rule le_imp_neg_le) - thus "a\b" by simp -next - assume "a\b" - thus "-b \ -a" by (rule le_imp_neg_le) -qed - -lemma neg_le_0_iff_le [simp]: "- a \ 0 \ 0 \ a" -by (subst neg_le_iff_le [symmetric], simp) - -lemma neg_0_le_iff_le [simp]: "0 \ - a \ a \ 0" -by (subst neg_le_iff_le [symmetric], simp) - -lemma neg_less_iff_less [simp]: "- b < - a \ a < b" -by (force simp add: less_le) - -lemma neg_less_0_iff_less [simp]: "- a < 0 \ 0 < a" -by (subst neg_less_iff_less [symmetric], simp) - -lemma neg_0_less_iff_less [simp]: "0 < - a \ a < 0" -by (subst neg_less_iff_less [symmetric], simp) - -text{*The next several equations can make the simplifier loop!*} - -lemma less_minus_iff: "a < - b \ b < - a" -proof - - have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less) - thus ?thesis by simp -qed - -lemma minus_less_iff: "- a < b \ - b < a" -proof - - have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less) - thus ?thesis by simp -qed - -lemma le_minus_iff: "a \ - b \ b \ - a" -proof - - have mm: "!! a (b::'a). (-(-a)) < -b \ -(-b) < -a" by (simp only: minus_less_iff) - have "(- (- a) <= -b) = (b <= - a)" - apply (auto simp only: le_less) - apply (drule mm) - apply (simp_all) - apply (drule mm[simplified], assumption) - done - then show ?thesis by simp -qed - -lemma minus_le_iff: "- a \ b \ - b \ a" -by (auto simp add: le_less minus_less_iff) - -lemma less_iff_diff_less_0: "a < b \ a - b < 0" -proof - - have "(a < b) = (a + (- b) < b + (-b))" - by (simp only: add_less_cancel_right) - also have "... = (a - b < 0)" by (simp add: diff_minus) - finally show ?thesis . -qed - -lemma diff_less_eq[algebra_simps]: "a - b < c \ a < c + b" -apply (subst less_iff_diff_less_0 [of a]) -apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst]) -apply (simp add: diff_minus add_ac) -done - -lemma less_diff_eq[algebra_simps]: "a < c - b \ a + b < c" -apply (subst less_iff_diff_less_0 [of "plus a b"]) -apply (subst less_iff_diff_less_0 [of a]) -apply (simp add: diff_minus add_ac) -done - -lemma diff_le_eq[algebra_simps]: "a - b \ c \ a \ c + b" -by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel) - -lemma le_diff_eq[algebra_simps]: "a \ c - b \ a + b \ c" -by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel) - -lemma le_iff_diff_le_0: "a \ b \ a - b \ 0" -by (simp add: algebra_simps) - -text{*Legacy - use @{text algebra_simps} *} -lemmas group_simps[noatp] = algebra_simps - -end - -text{*Legacy - use @{text algebra_simps} *} -lemmas group_simps[noatp] = algebra_simps - -class linordered_ab_semigroup_add = - linorder + ordered_ab_semigroup_add - -class linordered_cancel_ab_semigroup_add = - linorder + ordered_cancel_ab_semigroup_add -begin - -subclass linordered_ab_semigroup_add .. - -subclass ordered_ab_semigroup_add_imp_le -proof - fix a b c :: 'a - assume le: "c + a <= c + b" - show "a <= b" - proof (rule ccontr) - assume w: "~ a \ b" - hence "b <= a" by (simp add: linorder_not_le) - hence le2: "c + b <= c + a" by (rule add_left_mono) - have "a = b" - apply (insert le) - apply (insert le2) - apply (drule antisym, simp_all) - done - with w show False - by (simp add: linorder_not_le [symmetric]) - qed -qed - -end - -class linordered_ab_group_add = linorder + ordered_ab_group_add -begin - -subclass linordered_cancel_ab_semigroup_add .. - -lemma neg_less_eq_nonneg [simp]: - "- a \ a \ 0 \ a" -proof - assume A: "- a \ a" show "0 \ a" - proof (rule classical) - assume "\ 0 \ a" - then have "a < 0" by auto - with A have "- a < 0" by (rule le_less_trans) - then show ?thesis by auto - qed -next - assume A: "0 \ a" show "- a \ a" - proof (rule order_trans) - show "- a \ 0" using A by (simp add: minus_le_iff) - next - show "0 \ a" using A . - qed -qed - -lemma neg_less_nonneg [simp]: - "- a < a \ 0 < a" -proof - assume A: "- a < a" show "0 < a" - proof (rule classical) - assume "\ 0 < a" - then have "a \ 0" by auto - with A have "- a < 0" by (rule less_le_trans) - then show ?thesis by auto - qed -next - assume A: "0 < a" show "- a < a" - proof (rule less_trans) - show "- a < 0" using A by (simp add: minus_le_iff) - next - show "0 < a" using A . - qed -qed - -lemma less_eq_neg_nonpos [simp]: - "a \ - a \ a \ 0" -proof - assume A: "a \ - a" show "a \ 0" - proof (rule classical) - assume "\ a \ 0" - then have "0 < a" by auto - then have "0 < - a" using A by (rule less_le_trans) - then show ?thesis by auto - qed -next - assume A: "a \ 0" show "a \ - a" - proof (rule order_trans) - show "0 \ - a" using A by (simp add: minus_le_iff) - next - show "a \ 0" using A . - qed -qed - -lemma equal_neg_zero [simp]: - "a = - a \ a = 0" -proof - assume "a = 0" then show "a = - a" by simp -next - assume A: "a = - a" show "a = 0" - proof (cases "0 \ a") - case True with A have "0 \ - a" by auto - with le_minus_iff have "a \ 0" by simp - with True show ?thesis by (auto intro: order_trans) - next - case False then have B: "a \ 0" by auto - with A have "- a \ 0" by auto - with B show ?thesis by (auto intro: order_trans) - qed -qed - -lemma neg_equal_zero [simp]: - "- a = a \ a = 0" - by (auto dest: sym) - -lemma double_zero [simp]: - "a + a = 0 \ a = 0" -proof - assume assm: "a + a = 0" - then have a: "- a = a" by (rule minus_unique) - then show "a = 0" by (simp add: neg_equal_zero) -qed simp - -lemma double_zero_sym [simp]: - "0 = a + a \ a = 0" - by (rule, drule sym) simp_all - -lemma zero_less_double_add_iff_zero_less_single_add [simp]: - "0 < a + a \ 0 < a" -proof - assume "0 < a + a" - then have "0 - a < a" by (simp only: diff_less_eq) - then have "- a < a" by simp - then show "0 < a" by (simp add: neg_less_nonneg) -next - assume "0 < a" - with this have "0 + 0 < a + a" - by (rule add_strict_mono) - then show "0 < a + a" by simp -qed - -lemma zero_le_double_add_iff_zero_le_single_add [simp]: - "0 \ a + a \ 0 \ a" - by (auto simp add: le_less) - -lemma double_add_less_zero_iff_single_add_less_zero [simp]: - "a + a < 0 \ a < 0" -proof - - have "\ a + a < 0 \ \ a < 0" - by (simp add: not_less) - then show ?thesis by simp -qed - -lemma double_add_le_zero_iff_single_add_le_zero [simp]: - "a + a \ 0 \ a \ 0" -proof - - have "\ a + a \ 0 \ \ a \ 0" - by (simp add: not_le) - then show ?thesis by simp -qed - -lemma le_minus_self_iff: - "a \ - a \ a \ 0" -proof - - from add_le_cancel_left [of "- a" "a + a" 0] - have "a \ - a \ a + a \ 0" - by (simp add: add_assoc [symmetric]) - thus ?thesis by simp -qed - -lemma minus_le_self_iff: - "- a \ a \ 0 \ a" -proof - - from add_le_cancel_left [of "- a" 0 "a + a"] - have "- a \ a \ 0 \ a + a" - by (simp add: add_assoc [symmetric]) - thus ?thesis by simp -qed - -lemma minus_max_eq_min: - "- max x y = min (-x) (-y)" - by (auto simp add: max_def min_def) - -lemma minus_min_eq_max: - "- min x y = max (-x) (-y)" - by (auto simp add: max_def min_def) - -end - --- {* FIXME localize the following *} - -lemma add_increasing: - fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" - shows "[|0\a; b\c|] ==> b \ a + c" -by (insert add_mono [of 0 a b c], simp) - -lemma add_increasing2: - fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" - shows "[|0\c; b\a|] ==> b \ a + c" -by (simp add:add_increasing add_commute[of a]) - -lemma add_strict_increasing: - fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" - shows "[|0c|] ==> b < a + c" -by (insert add_less_le_mono [of 0 a b c], simp) - -lemma add_strict_increasing2: - fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}" - shows "[|0\a; b b < a + c" -by (insert add_le_less_mono [of 0 a b c], simp) - - -class ordered_ab_group_add_abs = ordered_ab_group_add + abs + - assumes abs_ge_zero [simp]: "\a\ \ 0" - and abs_ge_self: "a \ \a\" - and abs_leI: "a \ b \ - a \ b \ \a\ \ b" - and abs_minus_cancel [simp]: "\-a\ = \a\" - and abs_triangle_ineq: "\a + b\ \ \a\ + \b\" -begin - -lemma abs_minus_le_zero: "- \a\ \ 0" - unfolding neg_le_0_iff_le by simp - -lemma abs_of_nonneg [simp]: - assumes nonneg: "0 \ a" shows "\a\ = a" -proof (rule antisym) - from nonneg le_imp_neg_le have "- a \ 0" by simp - from this nonneg have "- a \ a" by (rule order_trans) - then show "\a\ \ a" by (auto intro: abs_leI) -qed (rule abs_ge_self) - -lemma abs_idempotent [simp]: "\\a\\ = \a\" -by (rule antisym) - (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"]) - -lemma abs_eq_0 [simp]: "\a\ = 0 \ a = 0" -proof - - have "\a\ = 0 \ a = 0" - proof (rule antisym) - assume zero: "\a\ = 0" - with abs_ge_self show "a \ 0" by auto - from zero have "\-a\ = 0" by simp - with abs_ge_self [of "uminus a"] have "- a \ 0" by auto - with neg_le_0_iff_le show "0 \ a" by auto - qed - then show ?thesis by auto -qed - -lemma abs_zero [simp]: "\0\ = 0" -by simp - -lemma abs_0_eq [simp, noatp]: "0 = \a\ \ a = 0" -proof - - have "0 = \a\ \ \a\ = 0" by (simp only: eq_ac) - thus ?thesis by simp -qed - -lemma abs_le_zero_iff [simp]: "\a\ \ 0 \ a = 0" -proof - assume "\a\ \ 0" - then have "\a\ = 0" by (rule antisym) simp - thus "a = 0" by simp -next - assume "a = 0" - thus "\a\ \ 0" by simp -qed - -lemma zero_less_abs_iff [simp]: "0 < \a\ \ a \ 0" -by (simp add: less_le) - -lemma abs_not_less_zero [simp]: "\ \a\ < 0" -proof - - have a: "\x y. x \ y \ \ y < x" by auto - show ?thesis by (simp add: a) -qed - -lemma abs_ge_minus_self: "- a \ \a\" -proof - - have "- a \ \-a\" by (rule abs_ge_self) - then show ?thesis by simp -qed - -lemma abs_minus_commute: - "\a - b\ = \b - a\" -proof - - have "\a - b\ = \- (a - b)\" by (simp only: abs_minus_cancel) - also have "... = \b - a\" by simp - finally show ?thesis . -qed - -lemma abs_of_pos: "0 < a \ \a\ = a" -by (rule abs_of_nonneg, rule less_imp_le) - -lemma abs_of_nonpos [simp]: - assumes "a \ 0" shows "\a\ = - a" -proof - - let ?b = "- a" - have "- ?b \ 0 \ \- ?b\ = - (- ?b)" - unfolding abs_minus_cancel [of "?b"] - unfolding neg_le_0_iff_le [of "?b"] - unfolding minus_minus by (erule abs_of_nonneg) - then show ?thesis using assms by auto -qed - -lemma abs_of_neg: "a < 0 \ \a\ = - a" -by (rule abs_of_nonpos, rule less_imp_le) - -lemma abs_le_D1: "\a\ \ b \ a \ b" -by (insert abs_ge_self, blast intro: order_trans) - -lemma abs_le_D2: "\a\ \ b \ - a \ b" -by (insert abs_le_D1 [of "uminus a"], simp) - -lemma abs_le_iff: "\a\ \ b \ a \ b \ - a \ b" -by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2) - -lemma abs_triangle_ineq2: "\a\ - \b\ \ \a - b\" - apply (simp add: algebra_simps) - apply (subgoal_tac "abs a = abs (plus b (minus a b))") - apply (erule ssubst) - apply (rule abs_triangle_ineq) - apply (rule arg_cong[of _ _ abs]) - apply (simp add: algebra_simps) -done - -lemma abs_triangle_ineq3: "\\a\ - \b\\ \ \a - b\" - apply (subst abs_le_iff) - apply auto - apply (rule abs_triangle_ineq2) - apply (subst abs_minus_commute) - apply (rule abs_triangle_ineq2) -done - -lemma abs_triangle_ineq4: "\a - b\ \ \a\ + \b\" -proof - - have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl) - also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq) - finally show ?thesis by simp -qed - -lemma abs_diff_triangle_ineq: "\a + b - (c + d)\ \ \a - c\ + \b - d\" -proof - - have "\a + b - (c+d)\ = \(a-c) + (b-d)\" by (simp add: diff_minus add_ac) - also have "... \ \a-c\ + \b-d\" by (rule abs_triangle_ineq) - finally show ?thesis . -qed - -lemma abs_add_abs [simp]: - "\\a\ + \b\\ = \a\ + \b\" (is "?L = ?R") -proof (rule antisym) - show "?L \ ?R" by(rule abs_ge_self) -next - have "?L \ \\a\\ + \\b\\" by(rule abs_triangle_ineq) - also have "\ = ?R" by simp - finally show "?L \ ?R" . -qed - -end - -text {* Needed for abelian cancellation simprocs: *} - -lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)" -apply (subst add_left_commute) -apply (subst add_left_cancel) -apply simp -done - -lemma add_cancel_end: "(x + (y + z) = y) = (x = - (z::'a::ab_group_add))" -apply (subst add_cancel_21[of _ _ _ 0, simplified]) -apply (simp add: add_right_cancel[symmetric, of "x" "-z" "z", simplified]) -done - -lemma less_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \ (x < y) = (x' < y')" -by (simp add: less_iff_diff_less_0[of x y] less_iff_diff_less_0[of x' y']) - -lemma le_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \ (y <= x) = (y' <= x')" -apply (simp add: le_iff_diff_le_0[of y x] le_iff_diff_le_0[of y' x']) -apply (simp add: neg_le_iff_le[symmetric, of "y-x" 0] neg_le_iff_le[symmetric, of "y'-x'" 0]) -done - -lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \ (x = y) = (x' = y')" -by (simp only: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y']) - -lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)" -by (simp add: diff_minus) - -lemma le_add_right_mono: - assumes - "a <= b + (c::'a::ordered_ab_group_add)" - "c <= d" - shows "a <= b + d" - apply (rule_tac order_trans[where y = "b+c"]) - apply (simp_all add: prems) - done - - -subsection {* Tools setup *} - -lemma add_mono_thms_linordered_semiring [noatp]: - fixes i j k :: "'a\ordered_ab_semigroup_add" - shows "i \ j \ k \ l \ i + k \ j + l" - and "i = j \ k \ l \ i + k \ j + l" - and "i \ j \ k = l \ i + k \ j + l" - and "i = j \ k = l \ i + k = j + l" -by (rule add_mono, clarify+)+ - -lemma add_mono_thms_linordered_field [noatp]: - fixes i j k :: "'a\ordered_cancel_ab_semigroup_add" - shows "i < j \ k = l \ i + k < j + l" - and "i = j \ k < l \ i + k < j + l" - and "i < j \ k \ l \ i + k < j + l" - and "i \ j \ k < l \ i + k < j + l" - and "i < j \ k < l \ i + k < j + l" -by (auto intro: add_strict_right_mono add_strict_left_mono - add_less_le_mono add_le_less_mono add_strict_mono) - -text{*Simplification of @{term "x-y < 0"}, etc.*} -lemmas diff_less_0_iff_less [simp, noatp] = less_iff_diff_less_0 [symmetric] -lemmas diff_le_0_iff_le [simp, noatp] = le_iff_diff_le_0 [symmetric] - -ML {* -structure ab_group_add_cancel = Abel_Cancel -( - -(* term order for abelian groups *) - -fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a') - [@{const_name Algebras.zero}, @{const_name Algebras.plus}, - @{const_name Algebras.uminus}, @{const_name Algebras.minus}] - | agrp_ord _ = ~1; - -fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS); - -local - val ac1 = mk_meta_eq @{thm add_assoc}; - val ac2 = mk_meta_eq @{thm add_commute}; - val ac3 = mk_meta_eq @{thm add_left_commute}; - fun solve_add_ac thy _ (_ $ (Const (@{const_name Algebras.plus},_) $ _ $ _) $ _) = - SOME ac1 - | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Algebras.plus},_) $ y $ z)) = - if termless_agrp (y, x) then SOME ac3 else NONE - | solve_add_ac thy _ (_ $ x $ y) = - if termless_agrp (y, x) then SOME ac2 else NONE - | solve_add_ac thy _ _ = NONE -in - val add_ac_proc = Simplifier.simproc @{theory} - "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac; -end; - -val eq_reflection = @{thm eq_reflection}; - -val T = @{typ "'a::ab_group_add"}; - -val cancel_ss = HOL_basic_ss settermless termless_agrp - addsimprocs [add_ac_proc] addsimps - [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_def}, - @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero}, - @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel}, - @{thm minus_add_cancel}]; - -val sum_pats = [@{cterm "x + y::'a::ab_group_add"}, @{cterm "x - y::'a::ab_group_add"}]; - -val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}]; - -val dest_eqI = - fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of; - -); -*} - -ML {* - Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv]; -*} - -code_modulename SML - OrderedGroup Arith - -code_modulename OCaml - OrderedGroup Arith - -code_modulename Haskell - OrderedGroup Arith - -end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/PReal.thy --- a/src/HOL/PReal.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/PReal.thy Mon Feb 08 17:12:38 2010 +0100 @@ -12,7 +12,7 @@ imports Rational begin -text{*Could be generalized and moved to @{text Ring_and_Field}*} +text{*Could be generalized and moved to @{text Groups}*} lemma add_eq_exists: "\x. a+x = (b::rat)" by (rule_tac x="b-a" in exI, simp) diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Presburger.thy Mon Feb 08 17:12:38 2010 +0100 @@ -30,8 +30,8 @@ "\(z ::'a::{linorder}).\x t) = True" "\(z ::'a::{linorder}).\x t) = False" "\(z ::'a::{linorder}).\x t) = False" - "\z.\(x::'a::{linorder,plus,Ring_and_Field.dvd})z.\(x::'a::{linorder,plus,Ring_and_Field.dvd}) d dvd x + s) = (\ d dvd x + s)" + "\z.\(x::'a::{linorder,plus,Rings.dvd})z.\(x::'a::{linorder,plus,Rings.dvd}) d dvd x + s) = (\ d dvd x + s)" "\z.\x(z ::'a::{linorder}).\x>z.(x \ t) = False" "\(z ::'a::{linorder}).\x>z.(x > t) = True" "\(z ::'a::{linorder}).\x>z.(x \ t) = True" - "\z.\(x::'a::{linorder,plus,Ring_and_Field.dvd})>z. (d dvd x + s) = (d dvd x + s)" - "\z.\(x::'a::{linorder,plus,Ring_and_Field.dvd})>z. (\ d dvd x + s) = (\ d dvd x + s)" + "\z.\(x::'a::{linorder,plus,Rings.dvd})>z. (d dvd x + s) = (d dvd x + s)" + "\z.\(x::'a::{linorder,plus,Rings.dvd})>z. (\ d dvd x + s) = (\ d dvd x + s)" "\z.\x>z. F = F" by ((erule exE, erule exE,rule_tac x="max z za" in exI,simp)+,(rule_tac x="t" in exI,fastsimp)+) simp_all @@ -56,8 +56,8 @@ \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" - "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \ \x k. (d dvd x + t) = (d dvd (x - k*D) + t)" - "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" + "(d::'a::{comm_ring,Rings.dvd}) dvd D \ \x k. (d dvd x + t) = (d dvd (x - k*D) + t)" + "(d::'a::{comm_ring,Rings.dvd}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" "\x k. F = F" apply (auto elim!: dvdE simp add: algebra_simps) unfolding mult_assoc [symmetric] left_distrib [symmetric] left_diff_distrib [symmetric] @@ -243,7 +243,7 @@ {fix x have "P x \ P (x - i * d)" using step.hyps by blast also have "\ \ P(x - (i + 1) * d)" using minus[THEN spec, of "x - i * d"] - by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric]) + by (simp add: algebra_simps) ultimately have "P x \ P(x - (i + 1) * d)" by blast} thus ?case .. qed @@ -360,7 +360,7 @@ apply(fastsimp) done -theorem unity_coeff_ex: "(\(x::'a::{semiring_0,Ring_and_Field.dvd}). P (l * x)) \ (\x. l dvd (x + 0) \ P x)" +theorem unity_coeff_ex: "(\(x::'a::{semiring_0,Rings.dvd}). P (l * x)) \ (\x. l dvd (x + 0) \ P x)" apply (rule eq_reflection [symmetric]) apply (rule iffI) defer diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Probability/Borel.thy --- a/src/HOL/Probability/Borel.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Probability/Borel.thy Mon Feb 08 17:12:38 2010 +0100 @@ -355,7 +355,7 @@ borel_measurable_add_borel_measurable f g) have "(\x. -((f x + -g x) ^ 2 * inverse 4)) = (\x. 0 + ((f x + -g x) ^ 2 * inverse -4))" - by (simp add: Ring_and_Field.minus_divide_right) + by (simp add: minus_divide_right) also have "... \ borel_measurable M" by (fast intro: affine_borel_measurable borel_measurable_square borel_measurable_add_borel_measurable diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/RealDef.thy Mon Feb 08 17:12:38 2010 +0100 @@ -521,7 +521,7 @@ lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" by (force elim: order_less_asym - simp add: Ring_and_Field.mult_less_cancel_right) + simp add: mult_less_cancel_right) lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \ y*z) = (x\y)" apply (simp add: mult_le_cancel_right) @@ -1015,7 +1015,7 @@ done -text{*Similar results are proved in @{text Ring_and_Field}*} +text{*Similar results are proved in @{text Fields}*} lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" by auto @@ -1030,7 +1030,7 @@ (* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *) lemma abs_le_interval_iff: "(abs x \ r) = (-r\x & x\(r::real))" -by (force simp add: OrderedGroup.abs_le_iff) +by (force simp add: abs_le_iff) lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)" by (simp add: abs_if) diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Mon Feb 08 17:12:32 2010 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2226 +0,0 @@ -(* Title: HOL/Ring_and_Field.thy - Author: Gertrud Bauer - Author: Steven Obua - Author: Tobias Nipkow - Author: Lawrence C Paulson - Author: Markus Wenzel - Author: Jeremy Avigad -*) - -header {* (Ordered) Rings and Fields *} - -theory Ring_and_Field -imports OrderedGroup -begin - -text {* - The theory of partially ordered rings is taken from the books: - \begin{itemize} - \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 - \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963 - \end{itemize} - Most of the used notions can also be looked up in - \begin{itemize} - \item \url{http://www.mathworld.com} by Eric Weisstein et. al. - \item \emph{Algebra I} by van der Waerden, Springer. - \end{itemize} -*} - -class semiring = ab_semigroup_add + semigroup_mult + - assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c" - assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c" -begin - -text{*For the @{text combine_numerals} simproc*} -lemma combine_common_factor: - "a * e + (b * e + c) = (a + b) * e + c" -by (simp add: left_distrib add_ac) - -end - -class mult_zero = times + zero + - assumes mult_zero_left [simp]: "0 * a = 0" - assumes mult_zero_right [simp]: "a * 0 = 0" - -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: left_distrib [symmetric]) - thus "0 * a = 0" by (simp only: add_left_cancel) -next - fix a :: 'a - have "a * 0 + a * 0 = a * 0 + 0" by (simp add: right_distrib [symmetric]) - thus "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: mult_ac) - also have "... = b * a + c * a" by (simp only: distrib) - also have "... = a * b + a * c" by (simp add: mult_ac) - 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) - -end - -class semiring_1 = zero_neq_one + semiring_0 + monoid_mult - -text {* Abstract divisibility *} - -class dvd = times -begin - -definition dvd :: "'a \ 'a \ bool" (infixl "dvd" 50) where - [code del]: "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 - -class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult + dvd - (*previously almost_semiring*) -begin - -subclass semiring_1 .. - -lemma dvd_refl[simp]: "a dvd a" -proof - show "a = a * 1" by simp -qed - -lemma dvd_trans: - assumes "a dvd b" and "b dvd c" - shows "a dvd c" -proof - - from assms obtain v where "b = a * v" by (auto elim!: dvdE) - moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE) - ultimately have "c = a * (v * w)" by (simp add: mult_assoc) - then show ?thesis .. -qed - -lemma dvd_0_left_iff [noatp, simp]: "0 dvd a \ a = 0" -by (auto intro: dvd_refl elim!: dvdE) - -lemma dvd_0_right [iff]: "a dvd 0" -proof - show "0 = a * 0" by simp -qed - -lemma one_dvd [simp]: "1 dvd a" -by (auto intro!: dvdI) - -lemma dvd_mult[simp]: "a dvd c \ a dvd (b * c)" -by (auto intro!: mult_left_commute dvdI elim!: dvdE) - -lemma dvd_mult2[simp]: "a dvd b \ a dvd (b * c)" - apply (subst mult_commute) - apply (erule dvd_mult) - done - -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: mult_ac) - 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" - unfolding mult_ac [of a] by (rule dvd_mult_left) - -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: right_distrib) - then show ?thesis .. -qed - -end - - -class no_zero_divisors = zero + times + - assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" - -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 -begin - -subclass semiring_1_cancel .. -subclass comm_semiring_0_cancel .. -subclass comm_semiring_1 .. - -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: left_distrib [symmetric]) - -lemma minus_mult_right: "- (a * b) = a * - b" -by (rule minus_unique) (simp add: right_distrib [symmetric]) - -text{*Extract signs from products*} -lemmas mult_minus_left [simp, noatp] = minus_mult_left [symmetric] -lemmas mult_minus_right [simp,noatp] = 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]: "a * (b - c) = a * b - a * c" -by (simp add: right_distrib diff_minus) - -lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c" -by (simp add: left_distrib diff_minus) - -lemmas ring_distribs[noatp] = - right_distrib left_distrib left_diff_distrib right_diff_distrib - -text{*Legacy - use @{text algebra_simps} *} -lemmas ring_simps[noatp] = algebra_simps - -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[noatp] = - right_distrib left_distrib left_diff_distrib right_diff_distrib - -class comm_ring = comm_semiring + ab_group_add -begin - -subclass ring .. -subclass comm_semiring_0_cancel .. - -end - -class ring_1 = ring + zero_neq_one + monoid_mult -begin - -subclass semiring_1_cancel .. - -end - -class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult - (*previously ring*) -begin - -subclass ring_1 .. -subclass comm_semiring_1_cancel .. - -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)" -by (simp add: diff_minus dvd_minus_iff) - -end - -class ring_no_zero_divisors = ring + no_zero_divisors -begin - -lemma mult_eq_0_iff [simp]: - shows "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 - -text{*Cancellation of equalities with a common factor*} -lemma mult_cancel_right [simp, noatp]: - "a * c = b * c \ c = 0 \ a = b" -proof - - have "(a * c = b * c) = ((a - b) * c = 0)" - by (simp add: algebra_simps right_minus_eq) - thus ?thesis by (simp add: disj_commute right_minus_eq) -qed - -lemma mult_cancel_left [simp, noatp]: - "c * a = c * b \ c = 0 \ a = b" -proof - - have "(c * a = c * b) = (c * (a - b) = 0)" - by (simp add: algebra_simps right_minus_eq) - thus ?thesis by (simp add: right_minus_eq) -qed - -end - -class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors -begin - -lemma mult_cancel_right1 [simp]: - "c = b * c \ c = 0 \ b = 1" -by (insert mult_cancel_right [of 1 c b], force) - -lemma mult_cancel_right2 [simp]: - "a * c = c \ c = 0 \ a = 1" -by (insert mult_cancel_right [of a c 1], simp) - -lemma mult_cancel_left1 [simp]: - "c = c * b \ c = 0 \ b = 1" -by (insert mult_cancel_left [of c 1 b], force) - -lemma mult_cancel_left2 [simp]: - "c * a = c \ c = 0 \ a = 1" -by (insert mult_cancel_left [of c a 1], simp) - -end - -class idom = comm_ring_1 + no_zero_divisors -begin - -subclass ring_1_no_zero_divisors .. - -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: right_minus_eq eq_neg_iff_add_eq_0) -next - assume "a = b \ a = - b" - then show "a * a = b * b" by auto -qed - -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)" - unfolding dvd_def by (simp add: mult_ac) - also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" - unfolding dvd_def by simp - finally show ?thesis . -qed - -lemma dvd_mult_cancel_left [simp]: - "c * a dvd c * b \ c = 0 \ a dvd b" -proof - - have "c * a dvd c * b \ (\k. b * c = (a * k) * c)" - unfolding dvd_def by (simp add: mult_ac) - also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" - unfolding dvd_def by simp - finally show ?thesis . -qed - -end - -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" -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: - "inverse a = 0 \ a = 0" -apply (rule classical) -apply (drule nonzero_imp_inverse_nonzero) -apply auto -done - -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) - -end - -class field = comm_ring_1 + inverse + - assumes field_inverse: "a \ 0 \ inverse a * a = 1" - assumes divide_inverse: "a / b = a * inverse b" -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) -qed - -subclass idom .. - -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_ac) - 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 divide_zero_left [simp]: "0 / a = 0" -by (simp add: divide_inverse) - -lemma inverse_eq_divide: "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) - -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 mult_ac) - -lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]: -assumes [simp]: "b\0" and [simp]: "c\0" shows "(c*a)/(c*b) = a/b" -proof - - 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: mult_ac) - also have "... = a * inverse b" by simp - finally show ?thesis by (simp add: divide_inverse) -qed - -lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]: - "\b \ 0; c \ 0\ \ (a * c) / (b * c) = a / b" -by (simp add: mult_commute [of _ c]) - -lemma divide_1 [simp]: "a / 1 = a" -by (simp add: divide_inverse) - -lemma times_divide_eq_right: "a * (b / c) = (a * b) / c" -by (simp add: divide_inverse mult_assoc) - -lemma times_divide_eq_left: "(b / c) * a = (b * a) / c" -by (simp add: divide_inverse mult_ac) - -text {* These are later declared as simp rules. *} -lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left - -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_mult_divide_cancel_right [simp, noatp]: - "b \ 0 \ a * b / b = a" -using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp - -lemma nonzero_mult_divide_cancel_left [simp, noatp]: - "a \ 0 \ a * b / a = b" -using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp - -lemma nonzero_divide_mult_cancel_right [simp, noatp]: - "\a \ 0; b \ 0\ \ b / (a * b) = 1 / a" -using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp - -lemma nonzero_divide_mult_cancel_left [simp, noatp]: - "\a \ 0; b \ 0\ \ a / (a * b) = 1 / b" -using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp - -lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]: - "\b \ 0; c \ 0\ \ (c * a) / (b * c) = a / b" -using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac) - -lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]: - "\b \ 0; c \ 0\ \ (a * c) / (c * b) = a / b" -using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac) - -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, noatp]: "(-a) / b = - (a / b)" -by (simp add: divide_inverse) - -lemma diff_divide_distrib: "(a - b) / c = a / c - b / c" -by (simp add: diff_minus add_divide_distrib) - -lemma add_divide_eq_iff: - "z \ 0 \ x + y / z = (z * x + y) / z" -by (simp add: add_divide_distrib) - -lemma divide_add_eq_iff: - "z \ 0 \ x / z + y = (x + z * y) / z" -by (simp add: add_divide_distrib) - -lemma diff_divide_eq_iff: - "z \ 0 \ x - y / z = (z * x - y) / z" -by (simp add: diff_divide_distrib) - -lemma divide_diff_eq_iff: - "z \ 0 \ x / z - y = (x - z * y) / z" -by (simp add: diff_divide_distrib) - -lemma nonzero_eq_divide_eq: "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: "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 divide_eq_imp: "c \ 0 \ b = a * c \ b / c = a" -by simp - -lemma eq_divide_imp: "c \ 0 \ a * c = b \ a = b / c" -by (erule subst, simp) - -lemmas field_eq_simps[noatp] = algebra_simps - (* pull / out*) - add_divide_eq_iff divide_add_eq_iff - diff_divide_eq_iff divide_diff_eq_iff - (* multiply eqn *) - nonzero_eq_divide_eq nonzero_divide_eq_eq -(* is added later: - times_divide_eq_left times_divide_eq_right -*) - -text{*An example:*} -lemma "\a\b; c\d; e\f\ \ ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1" -apply(subgoal_tac "(c-d)*(e-f)*(a-b) \ 0") - apply(simp add:field_eq_simps) -apply(simp) -done - -lemma diff_frac_eq: - "y \ 0 \ z \ 0 \ x / y - w / z = (x * z - w * y) / (y * z)" -by (simp add: field_eq_simps times_divide_eq) - -lemma frac_eq_eq: - "y \ 0 \ z \ 0 \ (x / y = w / z) = (x * z = w * y)" -by (simp add: field_eq_simps times_divide_eq) - -end - -class division_by_zero = zero + inverse + - assumes inverse_zero [simp]: "inverse 0 = 0" - -lemma divide_zero [simp]: - "a / 0 = (0::'a::{field,division_by_zero})" -by (simp add: divide_inverse) - -lemma divide_self_if [simp]: - "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)" -by simp - -class mult_mono = times + zero + ord + - assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" - assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" - -class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add -begin - -lemma mult_mono: - "a \ b \ c \ d \ 0 \ b \ 0 \ c - \ a * c \ b * d" -apply (erule mult_right_mono [THEN order_trans], assumption) -apply (erule mult_left_mono, assumption) -done - -lemma mult_mono': - "a \ b \ c \ d \ 0 \ a \ 0 \ c - \ a * c \ b * d" -apply (rule mult_mono) -apply (fast intro: order_trans)+ -done - -end - -class ordered_cancel_semiring = mult_mono + ordered_ab_semigroup_add - + semiring + cancel_comm_monoid_add -begin - -subclass semiring_0_cancel .. -subclass ordered_semiring .. - -lemma mult_nonneg_nonneg: "0 \ a \ 0 \ b \ 0 \ a * b" -using mult_left_mono [of zero b a] by simp - -lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" -using mult_left_mono [of b zero a] by simp - -lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" -using mult_right_mono [of a zero b] by simp - -text {* Legacy - use @{text mult_nonpos_nonneg} *} -lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" -by (drule mult_right_mono [of b zero], 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 linordered_semiring = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + mult_mono -begin - -subclass ordered_cancel_semiring .. - -subclass ordered_comm_monoid_add .. - -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 linordered_semiring_1 = linordered_semiring + semiring_1 - -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: "a \ b" "0 \ c" - from A show "c * a \ c * b" - unfolding le_less - using mult_strict_left_mono by (cases "c = 0") auto - from A 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 (force simp add: mult_strict_left_mono _not_less [symmetric]) - -lemma mult_right_le_imp_le: - "a * c \ b * c \ 0 < c \ a \ b" -by (force simp add: mult_strict_right_mono not_less [symmetric]) - -lemma mult_pos_pos: "0 < a \ 0 < b \ 0 < a * b" -using mult_strict_left_mono [of zero b a] by simp - -lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" -using mult_strict_left_mono [of b zero a] by simp - -lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" -using mult_strict_right_mono [of a zero b] by simp - -text {* Legacy - use @{text mult_neg_pos} *} -lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" -by (drule mult_strict_right_mono [of b zero], auto) - -lemma zero_less_mult_pos: - "0 < a * b \ 0 < a \ 0 < b" -apply (cases "b\0") - apply (auto simp add: le_less not_less) -apply (drule_tac mult_pos_neg [of a b]) - apply (auto dest: less_not_sym) -done - -lemma zero_less_mult_pos2: - "0 < b * a \ 0 < a \ 0 < b" -apply (cases "b\0") - apply (auto simp add: le_less not_less) -apply (drule_tac mult_pos_neg2 [of a b]) - apply (auto dest: less_not_sym) -done - -text{*Strict monotonicity in both arguments*} -lemma mult_strict_mono: - assumes "a < b" and "c < d" and "0 < b" and "0 \ c" - shows "a * c < b * d" - using assms apply (cases "c=0") - apply (simp add: mult_pos_pos) - apply (erule mult_strict_right_mono [THEN less_trans]) - apply (force simp add: le_less) - apply (erule mult_strict_left_mono, assumption) - done - -text{*This weaker variant has more natural premises*} -lemma mult_strict_mono': - assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" - shows "a * c < b * d" -by (rule mult_strict_mono) (insert assms, auto) - -lemma mult_less_le_imp_less: - assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" - shows "a * c < b * d" - using assms apply (subgoal_tac "a * c < b * c") - apply (erule less_le_trans) - apply (erule mult_left_mono) - apply simp - apply (erule mult_strict_right_mono) - apply assumption - done - -lemma mult_le_less_imp_less: - assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" - shows "a * c < b * d" - using assms apply (subgoal_tac "a * c \ b * c") - apply (erule le_less_trans) - apply (erule mult_strict_left_mono) - apply simp - apply (erule mult_right_mono) - apply simp - done - -lemma mult_less_imp_less_left: - assumes less: "c * a < c * b" and nonneg: "0 \ c" - shows "a < b" -proof (rule ccontr) - assume "\ a < b" - hence "b \ a" by (simp add: linorder_not_less) - hence "c * b \ c * a" using nonneg by (rule mult_left_mono) - with this and less show False by (simp add: not_less [symmetric]) -qed - -lemma mult_less_imp_less_right: - assumes less: "a * c < b * c" and nonneg: "0 \ c" - shows "a < b" -proof (rule ccontr) - assume "\ a < b" - hence "b \ a" by (simp add: linorder_not_less) - hence "b * c \ a * c" using nonneg by (rule mult_right_mono) - with this and less show False by (simp add: not_less [symmetric]) -qed - -end - -class linlinordered_semiring_1_strict = linordered_semiring_strict + semiring_1 - -class mult_mono1 = times + zero + ord + - assumes mult_mono1: "a \ b \ 0 \ c \ c * a \ c * b" - -class ordered_comm_semiring = comm_semiring_0 - + ordered_ab_semigroup_add + mult_mono1 -begin - -subclass ordered_semiring -proof - fix a b c :: 'a - assume "a \ b" "0 \ c" - thus "c * a \ c * b" by (rule mult_mono1) - thus "a * c \ b * c" by (simp only: mult_commute) -qed - -end - -class ordered_cancel_comm_semiring = comm_semiring_0_cancel - + ordered_ab_semigroup_add + mult_mono1 -begin - -subclass ordered_comm_semiring .. -subclass ordered_cancel_semiring .. - -end - -class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + - assumes mult_strict_left_mono_comm: "a < b \ 0 < c \ c * a < c * b" -begin - -subclass linordered_semiring_strict -proof - fix a b c :: 'a - assume "a < b" "0 < c" - thus "c * a < c * b" by (rule mult_strict_left_mono_comm) - thus "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" - thus "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 .. - -text{*Legacy - use @{text algebra_simps} *} -lemmas ring_simps[noatp] = algebra_simps - -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" - apply (drule mult_left_mono [of _ _ "uminus c"]) - apply (simp_all add: minus_mult_left [symmetric]) - done - -lemma mult_right_mono_neg: - "b \ a \ c \ 0 \ a * c \ b * c" - apply (drule mult_right_mono [of _ _ "uminus c"]) - apply (simp_all add: minus_mult_right [symmetric]) - done - -lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" -using mult_right_mono_neg [of a zero b] by simp - -lemma split_mult_pos_le: - "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" -by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos) - -end - -class abs_if = minus + uminus + ord + zero + abs + - assumes abs_if: "\a\ = (if a < 0 then - a else a)" - -class sgn_if = minus + uminus + zero + one + ord + sgn + - assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" - -lemma (in sgn_if) sgn0[simp]: "sgn 0 = 0" -by(simp add:sgn_if) - -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_less neg_less_eq_nonneg less_eq_neg_nonpos) - (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric] - neg_less_eq_nonneg less_eq_neg_nonpos, auto intro: add_nonneg_nonneg, - auto intro!: less_imp_le add_neg_neg) -qed (auto simp add: abs_if less_eq_neg_nonpos neg_equal_zero) - -end - -(* The "strict" suffix can be seen as describing the combination of linordered_ring and no_zero_divisors. - Basically, linordered_ring + no_zero_divisors = linordered_ring_strict. - *) -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 zero 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 note A' = this - show ?thesis proof (cases "b < 0") - case True with A' - show ?thesis by (auto dest: mult_neg_neg) - next - case False with B have "0 < b" by auto - with A' show ?thesis by (auto dest: mult_strict_right_mono) - qed - next - case False with A have A': "0 < a" by auto - show ?thesis proof (cases "b < 0") - case True with A' - show ?thesis by (auto dest: mult_strict_right_mono_neg) - next - case False with B have "0 < b" by auto - with A' show ?thesis by (auto dest: mult_pos_pos) - qed - qed - then show "a * b \ 0" by (simp add: neq_iff) -qed - -lemma zero_less_mult_iff: - "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" - apply (auto simp add: mult_pos_pos mult_neg_neg) - apply (simp_all add: not_less le_less) - apply (erule disjE) apply assumption defer - apply (erule disjE) defer apply (drule sym) apply simp - apply (erule disjE) defer apply (drule sym) apply simp - apply (erule disjE) apply assumption apply (drule sym) apply simp - apply (drule sym) apply simp - apply (blast dest: zero_less_mult_pos) - apply (blast dest: zero_less_mult_pos2) - done - -lemma zero_le_mult_iff: - "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: - "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" - apply (insert zero_less_mult_iff [of "-a" b]) - apply (force simp add: minus_mult_left[symmetric]) - done - -lemma mult_le_0_iff: - "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" - apply (insert zero_le_mult_iff [of "-a" b]) - apply (force simp add: minus_mult_left[symmetric]) - done - -lemma zero_le_square [simp]: "0 \ a * a" -by (simp add: zero_le_mult_iff linear) - -lemma not_square_less_zero [simp]: "\ (a * a < 0)" -by (simp add: not_less) - -text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"}, - also with the relations @{text "\"} 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" - apply (cases "c = 0") - apply (auto simp add: neq_iff mult_strict_right_mono - mult_strict_right_mono_neg) - apply (auto simp add: not_less - not_le [symmetric, of "a*c"] - not_le [symmetric, of a]) - apply (erule_tac [!] notE) - apply (auto simp add: less_imp_le mult_right_mono - mult_right_mono_neg) - done - -lemma mult_less_cancel_left_disj: - "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" - apply (cases "c = 0") - apply (auto simp add: neq_iff mult_strict_left_mono - mult_strict_left_mono_neg) - apply (auto simp add: not_less - not_le [symmetric, of "c*a"] - not_le [symmetric, of a]) - apply (erule_tac [!] notE) - apply (auto simp add: less_imp_le mult_left_mono - mult_left_mono_neg) - done - -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 - -text{*Legacy - use @{text algebra_simps} *} -lemmas ring_simps[noatp] = algebra_simps - -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_semidom = comm_semiring_1_cancel + linordered_comm_semiring_strict + - (*previously linordered_semiring*) - assumes zero_less_one [simp]: "0 < 1" -begin - -lemma pos_add_strict: - shows "0 < a \ b < c \ b < a + c" - using add_strict_mono [of zero a b c] by simp - -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 less_1_mult: - assumes "1 < m" and "1 < n" - shows "1 < m * n" - using assms 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_if - (*previously linordered_ring*) -begin - -subclass linordered_ring_strict .. -subclass ordered_comm_ring .. -subclass idom .. - -subclass linordered_semidom -proof - have "0 \ 1 * 1" by (rule zero_le_square) - thus "0 < 1" by (simp add: le_less) -qed - -lemma linorder_neqE_linordered_idom: - assumes "x \ y" obtains "x < y" | "y < x" - using assms by (rule neqE) - -text {* These cancellation simprules 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)" -by (insert mult_le_cancel_right [of 1 c b], simp) - -lemma mult_le_cancel_right2: - "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" -by (insert mult_le_cancel_right [of a c 1], simp) - -lemma mult_le_cancel_left1: - "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" -by (insert mult_le_cancel_left [of c 1 b], simp) - -lemma mult_le_cancel_left2: - "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" -by (insert mult_le_cancel_left [of c a 1], simp) - -lemma mult_less_cancel_right1: - "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" -by (insert mult_less_cancel_right [of 1 c b], simp) - -lemma mult_less_cancel_right2: - "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" -by (insert mult_less_cancel_right [of a c 1], simp) - -lemma mult_less_cancel_left1: - "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" -by (insert mult_less_cancel_left [of c 1 b], simp) - -lemma mult_less_cancel_left2: - "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" -by (insert mult_less_cancel_left [of c a 1], simp) - -lemma sgn_sgn [simp]: - "sgn (sgn a) = sgn a" -unfolding sgn_if by simp - -lemma sgn_0_0: - "sgn a = 0 \ a = 0" -unfolding sgn_if by simp - -lemma sgn_1_pos: - "sgn a = 1 \ a > 0" -unfolding sgn_if by (simp add: neg_equal_zero) - -lemma sgn_1_neg: - "sgn a = - 1 \ a < 0" -unfolding sgn_if by (auto simp add: equal_neg_zero) - -lemma sgn_pos [simp]: - "0 < a \ sgn a = 1" -unfolding sgn_1_pos . - -lemma sgn_neg [simp]: - "a < 0 \ sgn a = - 1" -unfolding sgn_1_neg . - -lemma sgn_times: - "sgn (a * b) = sgn a * sgn b" -by (auto simp add: sgn_if zero_less_mult_iff) - -lemma abs_sgn: "abs 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_dvd_iff [simp]: "(abs m) dvd k \ m dvd k" - by (simp add: abs_if) - -lemma dvd_abs_iff [simp]: "m dvd (abs k) \ m dvd k" - by (simp add: abs_if) - -lemma dvd_if_abs_eq: - "abs l = abs (k) \ l dvd k" -by(subst abs_dvd_iff[symmetric]) simp - -end - -class linordered_field = field + linordered_idom - -text {* Simprules for comparisons where common factors can be cancelled. *} - -lemmas mult_compare_simps[noatp] = - mult_le_cancel_right mult_le_cancel_left - mult_le_cancel_right1 mult_le_cancel_right2 - mult_le_cancel_left1 mult_le_cancel_left2 - mult_less_cancel_right mult_less_cancel_left - mult_less_cancel_right1 mult_less_cancel_right2 - mult_less_cancel_left1 mult_less_cancel_left2 - mult_cancel_right mult_cancel_left - mult_cancel_right1 mult_cancel_right2 - mult_cancel_left1 mult_cancel_left2 - --- {* FIXME continue localization here *} - -lemma inverse_nonzero_iff_nonzero [simp]: - "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))" -by (force dest: inverse_zero_imp_zero) - -lemma inverse_minus_eq [simp]: - "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})" -proof cases - assume "a=0" thus ?thesis by (simp add: inverse_zero) -next - assume "a\0" - thus ?thesis by (simp add: nonzero_inverse_minus_eq) -qed - -lemma inverse_eq_imp_eq: - "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})" -apply (cases "a=0 | b=0") - apply (force dest!: inverse_zero_imp_zero - simp add: eq_commute [of "0::'a"]) -apply (force dest!: nonzero_inverse_eq_imp_eq) -done - -lemma inverse_eq_iff_eq [simp]: - "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))" -by (force dest!: inverse_eq_imp_eq) - -lemma inverse_inverse_eq [simp]: - "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a" - proof cases - assume "a=0" thus ?thesis by simp - next - assume "a\0" - thus ?thesis by (simp add: nonzero_inverse_inverse_eq) - qed - -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::'a::{field,division_by_zero})" - proof cases - assume "a \ 0 & b \ 0" - thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute) - next - assume "~ (a \ 0 & b \ 0)" - thus ?thesis by force - qed - -lemma inverse_divide [simp]: - "inverse (a/b) = b / (a::'a::{field,division_by_zero})" -by (simp add: divide_inverse mult_commute) - - -subsection {* Calculations with fractions *} - -text{* There is a whole bunch of simp-rules just for class @{text -field} but none for class @{text field} and @{text nonzero_divides} -because the latter are covered by a simproc. *} - -lemma mult_divide_mult_cancel_left: - "c\0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})" -apply (cases "b = 0") -apply (simp_all add: nonzero_mult_divide_mult_cancel_left) -done - -lemma mult_divide_mult_cancel_right: - "c\0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})" -apply (cases "b = 0") -apply (simp_all add: nonzero_mult_divide_mult_cancel_right) -done - -lemma divide_divide_eq_right [simp,noatp]: - "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})" -by (simp add: divide_inverse mult_ac) - -lemma divide_divide_eq_left [simp,noatp]: - "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)" -by (simp add: divide_inverse mult_assoc) - - -subsubsection{*Special Cancellation Simprules for Division*} - -lemma mult_divide_mult_cancel_left_if[simp,noatp]: -fixes c :: "'a :: {field,division_by_zero}" -shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)" -by (simp add: mult_divide_mult_cancel_left) - - -subsection {* Division and Unary Minus *} - -lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})" -by (simp add: divide_inverse) - -lemma divide_minus_right [simp, noatp]: - "a / -(b::'a::{field,division_by_zero}) = -(a / b)" -by (simp add: divide_inverse) - -lemma minus_divide_divide: - "(-a)/(-b) = a / (b::'a::{field,division_by_zero})" -apply (cases "b=0", simp) -apply (simp add: nonzero_minus_divide_divide) -done - -lemma eq_divide_eq: - "((a::'a::{field,division_by_zero}) = b/c) = (if c\0 then a*c = b else a=0)" -by (simp add: nonzero_eq_divide_eq) - -lemma divide_eq_eq: - "(b/c = (a::'a::{field,division_by_zero})) = (if c\0 then b = a*c else a=0)" -by (force simp add: nonzero_divide_eq_eq) - - -subsection {* Ordered Fields *} - -lemma positive_imp_inverse_positive: -assumes a_gt_0: "0 < a" shows "0 < inverse (a::'a::linordered_field)" -proof - - have "0 < a * inverse a" - by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one) - thus "0 < inverse a" - by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff) -qed - -lemma negative_imp_inverse_negative: - "a < 0 ==> inverse a < (0::'a::linordered_field)" -by (insert positive_imp_inverse_positive [of "-a"], - simp add: nonzero_inverse_minus_eq order_less_imp_not_eq) - -lemma inverse_le_imp_le: -assumes invle: "inverse a \ inverse b" and apos: "0 < a" -shows "b \ (a::'a::linordered_field)" -proof (rule classical) - assume "~ b \ a" - hence "a < b" by (simp add: linorder_not_le) - hence bpos: "0 < b" by (blast intro: apos order_less_trans) - hence "a * inverse a \ a * inverse b" - by (simp add: apos invle order_less_imp_le mult_left_mono) - hence "(a * inverse a) * b \ (a * inverse b) * b" - by (simp add: bpos order_less_imp_le mult_right_mono) - thus "b \ a" by (simp add: mult_assoc apos bpos order_less_imp_not_eq2) -qed - -lemma inverse_positive_imp_positive: -assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" -shows "0 < (a::'a::linordered_field)" -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_positive_iff_positive [simp]: - "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))" -apply (cases "a = 0", simp) -apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive) -done - -lemma inverse_negative_imp_negative: -assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" -shows "a < (0::'a::linordered_field)" -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 inverse_negative_iff_negative [simp]: - "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))" -apply (cases "a = 0", simp) -apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative) -done - -lemma inverse_nonnegative_iff_nonnegative [simp]: - "(0 \ inverse a) = (0 \ (a::'a::{linordered_field,division_by_zero}))" -by (simp add: linorder_not_less [symmetric]) - -lemma inverse_nonpositive_iff_nonpositive [simp]: - "(inverse a \ 0) = (a \ (0::'a::{linordered_field,division_by_zero}))" -by (simp add: linorder_not_less [symmetric]) - -lemma linordered_field_no_lb: "\ x. \y. y < (x::'a::linordered_field)" -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::'a::linordered_field)" -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 - -subsection{*Anti-Monotonicity of @{term inverse}*} - -lemma less_imp_inverse_less: -assumes less: "a < b" and apos: "0 < a" -shows "inverse b < inverse (a::'a::linordered_field)" -proof (rule ccontr) - assume "~ inverse b < inverse a" - hence "inverse a \ inverse b" by (simp add: linorder_not_less) - hence "~ (a < b)" - by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos]) - thus False by (rule notE [OF _ less]) -qed - -lemma inverse_less_imp_less: - "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)" -apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"]) -apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) -done - -text{*Both premises are essential. Consider -1 and 1.*} -lemma inverse_less_iff_less [simp,noatp]: - "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))" -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::'a::linordered_field)" -by (force simp add: order_le_less less_imp_inverse_less) - -lemma inverse_le_iff_le [simp,noatp]: - "[|0 < a; 0 < b|] ==> (inverse a \ inverse b) = (b \ (a::'a::linordered_field))" -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: - "[|inverse a \ inverse b; b < 0|] ==> b \ (a::'a::linordered_field)" -apply (rule classical) -apply (subgoal_tac "a < 0") - prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) -apply (insert inverse_le_imp_le [of "-b" "-a"]) -apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) -done - -lemma less_imp_inverse_less_neg: - "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)" -apply (subgoal_tac "a < 0") - prefer 2 apply (blast intro: order_less_trans) -apply (insert less_imp_inverse_less [of "-b" "-a"]) -apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) -done - -lemma inverse_less_imp_less_neg: - "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)" -apply (rule classical) -apply (subgoal_tac "a < 0") - prefer 2 - apply (force simp add: linorder_not_less intro: order_le_less_trans) -apply (insert inverse_less_imp_less [of "-b" "-a"]) -apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) -done - -lemma inverse_less_iff_less_neg [simp,noatp]: - "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))" -apply (insert inverse_less_iff_less [of "-b" "-a"]) -apply (simp del: inverse_less_iff_less - add: order_less_imp_not_eq nonzero_inverse_minus_eq) -done - -lemma le_imp_inverse_le_neg: - "[|a \ b; b < 0|] ==> inverse b \ inverse (a::'a::linordered_field)" -by (force simp add: order_le_less less_imp_inverse_less_neg) - -lemma inverse_le_iff_le_neg [simp,noatp]: - "[|a < 0; b < 0|] ==> (inverse a \ inverse b) = (b \ (a::'a::linordered_field))" -by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) - - -subsection{*Inverses and the Number One*} - -lemma one_less_inverse_iff: - "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))" -proof cases - assume "0 < x" - with inverse_less_iff_less [OF zero_less_one, of x] - show ?thesis by simp -next - assume notless: "~ (0 < x)" - have "~ (1 < inverse x)" - proof - assume "1 < inverse x" - also with notless have "... \ 0" by (simp add: linorder_not_less) - also have "... < 1" by (rule zero_less_one) - finally show False by auto - qed - with notless show ?thesis by simp -qed - -lemma inverse_eq_1_iff [simp]: - "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))" -by (insert inverse_eq_iff_eq [of x 1], simp) - -lemma one_le_inverse_iff: - "(1 \ inverse x) = (0 < x & x \ (1::'a::{linordered_field,division_by_zero}))" -by (force simp add: order_le_less one_less_inverse_iff zero_less_one - eq_commute [of 1]) - -lemma inverse_less_1_iff: - "(inverse x < 1) = (x \ 0 | 1 < (x::'a::{linordered_field,division_by_zero}))" -by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) - -lemma inverse_le_1_iff: - "(inverse x \ 1) = (x \ 0 | 1 \ (x::'a::{linordered_field,division_by_zero}))" -by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) - - -subsection{*Simplification of Inequalities Involving Literal Divisors*} - -lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \ b/c) = (a*c \ b)" -proof - - assume less: "0 b/c) = (a*c \ (b/c)*c)" - by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) - also have "... = (a*c \ b)" - by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \ b/c) = (b \ a*c)" -proof - - assume less: "c<0" - hence "(a \ b/c) = ((b/c)*c \ a*c)" - by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) - also have "... = (b \ a*c)" - by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma le_divide_eq: - "(a \ b/c) = - (if 0 < c then a*c \ b - else if c < 0 then b \ a*c - else a \ (0::'a::{linordered_field,division_by_zero}))" -apply (cases "c=0", simp) -apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) -done - -lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \ a) = (b \ a*c)" -proof - - assume less: "0 a) = ((b/c)*c \ a*c)" - by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) - also have "... = (b \ a*c)" - by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \ a) = (a*c \ b)" -proof - - assume less: "c<0" - hence "(b/c \ a) = (a*c \ (b/c)*c)" - by (simp add: mult_le_cancel_right order_less_not_sym [OF less]) - also have "... = (a*c \ b)" - by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma divide_le_eq: - "(b/c \ a) = - (if 0 < c then b \ a*c - else if c < 0 then a*c \ b - else 0 \ (a::'a::{linordered_field,division_by_zero}))" -apply (cases "c=0", simp) -apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) -done - -lemma pos_less_divide_eq: - "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)" -proof - - assume less: "0 (a < b/c) = (b < a*c)" -proof - - assume less: "c<0" - hence "(a < b/c) = ((b/c)*c < a*c)" - by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) - also have "... = (b < a*c)" - by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma less_divide_eq: - "(a < b/c) = - (if 0 < c then a*c < b - else if c < 0 then b < a*c - else a < (0::'a::{linordered_field,division_by_zero}))" -apply (cases "c=0", simp) -apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) -done - -lemma pos_divide_less_eq: - "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)" -proof - - assume less: "0 (b/c < a) = (a*c < b)" -proof - - assume less: "c<0" - hence "(b/c < a) = (a*c < (b/c)*c)" - by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less]) - also have "... = (a*c < b)" - by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) - finally show ?thesis . -qed - -lemma divide_less_eq: - "(b/c < a) = - (if 0 < c then b < a*c - else if c < 0 then a*c < b - else 0 < (a::'a::{linordered_field,division_by_zero}))" -apply (cases "c=0", simp) -apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) -done - - -subsection{*Field simplification*} - -text{* Lemmas @{text field_simps} multiply with denominators in in(equations) -if they can be proved to be non-zero (for equations) or positive/negative -(for inequations). Can be too aggressive and is therefore separate from the -more benign @{text algebra_simps}. *} - -lemmas field_simps[noatp] = field_eq_simps - (* multiply ineqn *) - pos_divide_less_eq neg_divide_less_eq - pos_less_divide_eq neg_less_divide_eq - pos_divide_le_eq neg_divide_le_eq - pos_le_divide_eq neg_le_divide_eq - -text{* Lemmas @{text sign_simps} is a first attempt to automate proofs -of positivity/negativity needed for @{text field_simps}. Have not added @{text -sign_simps} to @{text field_simps} because the former can lead to case -explosions. *} - -lemmas sign_simps[noatp] = group_simps - zero_less_mult_iff mult_less_0_iff - -(* Only works once linear arithmetic is installed: -text{*An example:*} -lemma fixes a b c d e f :: "'a::linordered_field" -shows "\a>b; c \ - ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) < - ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u" -apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0") - prefer 2 apply(simp add:sign_simps) -apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0") - prefer 2 apply(simp add:sign_simps) -apply(simp add:field_simps) -done -*) - - -subsection{*Division and Signs*} - -lemma zero_less_divide_iff: - "((0::'a::{linordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)" -by (simp add: divide_inverse zero_less_mult_iff) - -lemma divide_less_0_iff: - "(a/b < (0::'a::{linordered_field,division_by_zero})) = - (0 < a & b < 0 | a < 0 & 0 < b)" -by (simp add: divide_inverse mult_less_0_iff) - -lemma zero_le_divide_iff: - "((0::'a::{linordered_field,division_by_zero}) \ a/b) = - (0 \ a & 0 \ b | a \ 0 & b \ 0)" -by (simp add: divide_inverse zero_le_mult_iff) - -lemma divide_le_0_iff: - "(a/b \ (0::'a::{linordered_field,division_by_zero})) = - (0 \ a & b \ 0 | a \ 0 & 0 \ b)" -by (simp add: divide_inverse mult_le_0_iff) - -lemma divide_eq_0_iff [simp,noatp]: - "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))" -by (simp add: divide_inverse) - -lemma divide_pos_pos: - "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y" -by(simp add:field_simps) - - -lemma divide_nonneg_pos: - "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y" -by(simp add:field_simps) - -lemma divide_neg_pos: - "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0" -by(simp add:field_simps) - -lemma divide_nonpos_pos: - "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0" -by(simp add:field_simps) - -lemma divide_pos_neg: - "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0" -by(simp add:field_simps) - -lemma divide_nonneg_neg: - "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0" -by(simp add:field_simps) - -lemma divide_neg_neg: - "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y" -by(simp add:field_simps) - -lemma divide_nonpos_neg: - "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y" -by(simp add:field_simps) - - -subsection{*Cancellation Laws for Division*} - -lemma divide_cancel_right [simp,noatp]: - "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))" -apply (cases "c=0", simp) -apply (simp add: divide_inverse) -done - -lemma divide_cancel_left [simp,noatp]: - "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" -apply (cases "c=0", simp) -apply (simp add: divide_inverse) -done - - -subsection {* Division and the Number One *} - -text{*Simplify expressions equated with 1*} -lemma divide_eq_1_iff [simp,noatp]: - "(a/b = 1) = (b \ 0 & a = (b::'a::{field,division_by_zero}))" -apply (cases "b=0", simp) -apply (simp add: right_inverse_eq) -done - -lemma one_eq_divide_iff [simp,noatp]: - "(1 = a/b) = (b \ 0 & a = (b::'a::{field,division_by_zero}))" -by (simp add: eq_commute [of 1]) - -lemma zero_eq_1_divide_iff [simp,noatp]: - "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)" -apply (cases "a=0", simp) -apply (auto simp add: nonzero_eq_divide_eq) -done - -lemma one_divide_eq_0_iff [simp,noatp]: - "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)" -apply (cases "a=0", simp) -apply (insert zero_neq_one [THEN not_sym]) -apply (auto simp add: nonzero_divide_eq_eq) -done - -text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*} -lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified] -lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified] -lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified] -lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified] - -declare zero_less_divide_1_iff [simp,noatp] -declare divide_less_0_1_iff [simp,noatp] -declare zero_le_divide_1_iff [simp,noatp] -declare divide_le_0_1_iff [simp,noatp] - - -subsection {* Ordering Rules for Division *} - -lemma divide_strict_right_mono: - "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)" -by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono - positive_imp_inverse_positive) - -lemma divide_right_mono: - "[|a \ b; 0 \ c|] ==> a/c \ b/(c::'a::{linordered_field,division_by_zero})" -by (force simp add: divide_strict_right_mono order_le_less) - -lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b - ==> c <= 0 ==> b / c <= a / c" -apply (drule divide_right_mono [of _ _ "- c"]) -apply auto -done - -lemma divide_strict_right_mono_neg: - "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)" -apply (drule divide_strict_right_mono [of _ _ "-c"], simp) -apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric]) -done - -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::'a::linordered_field)" -by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono) - -lemma divide_left_mono: - "[|b \ a; 0 \ c; 0 < a*b|] ==> c / a \ c / (b::'a::linordered_field)" -by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono) - -lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b - ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b" - apply (drule divide_left_mono [of _ _ "- c"]) - apply (auto simp add: mult_commute) -done - -lemma divide_strict_left_mono_neg: - "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)" -by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg) - - -text{*Simplify quotients that are compared with the value 1.*} - -lemma le_divide_eq_1 [noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(1 \ b / a) = ((0 < a & a \ b) | (a < 0 & b \ a))" -by (auto simp add: le_divide_eq) - -lemma divide_le_eq_1 [noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(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 [noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))" -by (auto simp add: less_divide_eq) - -lemma divide_less_eq_1 [noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)" -by (auto simp add: divide_less_eq) - - -subsection{*Conditional Simplification Rules: No Case Splits*} - -lemma le_divide_eq_1_pos [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "0 < a \ (1 \ b/a) = (a \ b)" -by (auto simp add: le_divide_eq) - -lemma le_divide_eq_1_neg [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "a < 0 \ (1 \ b/a) = (b \ a)" -by (auto simp add: le_divide_eq) - -lemma divide_le_eq_1_pos [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "0 < a \ (b/a \ 1) = (b \ a)" -by (auto simp add: divide_le_eq) - -lemma divide_le_eq_1_neg [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "a < 0 \ (b/a \ 1) = (a \ b)" -by (auto simp add: divide_le_eq) - -lemma less_divide_eq_1_pos [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "0 < a \ (1 < b/a) = (a < b)" -by (auto simp add: less_divide_eq) - -lemma less_divide_eq_1_neg [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "a < 0 \ (1 < b/a) = (b < a)" -by (auto simp add: less_divide_eq) - -lemma divide_less_eq_1_pos [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "0 < a \ (b/a < 1) = (b < a)" -by (auto simp add: divide_less_eq) - -lemma divide_less_eq_1_neg [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "a < 0 \ b/a < 1 <-> a < b" -by (auto simp add: divide_less_eq) - -lemma eq_divide_eq_1 [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(1 = b/a) = ((a \ 0 & a = b))" -by (auto simp add: eq_divide_eq) - -lemma divide_eq_eq_1 [simp,noatp]: - fixes a :: "'a :: {linordered_field,division_by_zero}" - shows "(b/a = 1) = ((a \ 0 & a = b))" -by (auto simp add: divide_eq_eq) - - -subsection {* Reasoning about inequalities with division *} - -lemma mult_right_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1 - ==> x * y <= x" -by (auto simp add: mult_compare_simps) - -lemma mult_left_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1 - ==> y * x <= x" -by (auto simp add: mult_compare_simps) - -lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==> - x / y <= z" -by (subst pos_divide_le_eq, assumption+) - -lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==> - z <= x / y" -by(simp add:field_simps) - -lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==> - x / y < z" -by(simp add:field_simps) - -lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==> - z < x / y" -by(simp add:field_simps) - -lemma frac_le: "(0::'a::linordered_field) <= x ==> - x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w" - apply (rule mult_imp_div_pos_le) - apply simp - apply (subst times_divide_eq_left) - apply (rule mult_imp_le_div_pos, assumption) - apply (rule mult_mono) - apply simp_all -done - -lemma frac_less: "(0::'a::linordered_field) <= x ==> - x < y ==> 0 < w ==> w <= z ==> x / z < y / w" - apply (rule mult_imp_div_pos_less) - apply simp - apply (subst times_divide_eq_left) - apply (rule mult_imp_less_div_pos, assumption) - apply (erule mult_less_le_imp_less) - apply simp_all -done - -lemma frac_less2: "(0::'a::linordered_field) < x ==> - x <= y ==> 0 < w ==> w < z ==> x / z < y / w" - apply (rule mult_imp_div_pos_less) - apply simp_all - apply (subst times_divide_eq_left) - apply (rule mult_imp_less_div_pos, assumption) - apply (erule mult_le_less_imp_less) - apply simp_all -done - -text{*It's not obvious whether these should be simprules or not. - Their effect is to gather terms into one big fraction, like - a*b*c / x*y*z. The rationale for that is unclear, but many proofs - seem to need them.*} - -declare times_divide_eq [simp] - - -subsection {* Ordered Fields are Dense *} - -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) - thus ?thesis by simp -qed - -lemma zero_less_two: "0 < 1 + 1" -by (blast intro: less_trans zero_less_one less_add_one) - -end - -lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)" -by (simp add: field_simps zero_less_two) - -lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b" -by (simp add: field_simps zero_less_two) - -instance linordered_field < dense_linorder -proof - fix x y :: 'a - have "x < x + 1" by simp - then show "\y. x < y" .. - have "x - 1 < x" by simp - then show "\y. y < x" .. - show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) -qed - - -subsection {* Absolute Value *} - -context linordered_idom -begin - -lemma mult_sgn_abs: "sgn x * abs x = x" - unfolding abs_if sgn_if by auto - -end - -lemma abs_one [simp]: "abs 1 = (1::'a::linordered_idom)" -by (simp add: abs_if zero_less_one [THEN order_less_not_sym]) - -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 proof -qed (auto simp add: abs_if not_less equal_neg_zero neg_equal_zero mult_less_0_iff) - -lemma abs_mult: - "abs (a * b) = abs a * abs b" - by (rule abs_eq_mult) auto - -lemma abs_mult_self: - "abs a * abs a = a * a" - by (simp add: abs_if) - -end - -lemma nonzero_abs_inverse: - "a \ 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)" -apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq - negative_imp_inverse_negative) -apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) -done - -lemma abs_inverse [simp]: - "abs (inverse (a::'a::{linordered_field,division_by_zero})) = - inverse (abs a)" -apply (cases "a=0", simp) -apply (simp add: nonzero_abs_inverse) -done - -lemma nonzero_abs_divide: - "b \ 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b" -by (simp add: divide_inverse abs_mult nonzero_abs_inverse) - -lemma abs_divide [simp]: - "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b" -apply (cases "b=0", simp) -apply (simp add: nonzero_abs_divide) -done - -lemma abs_mult_less: - "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::linordered_idom)" -proof - - assume ac: "abs a < c" - hence cpos: "0 - (abs y) * x = abs (y * x)" - apply (subst abs_mult) - apply simp -done - -lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==> - abs x / y = abs (x / y)" - apply (subst abs_divide) - apply (simp add: order_less_imp_le) -done - -code_modulename SML - Ring_and_Field Arith - -code_modulename OCaml - Ring_and_Field Arith - -code_modulename Haskell - Ring_and_Field Arith - -end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Rings.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Rings.thy Mon Feb 08 17:12:38 2010 +0100 @@ -0,0 +1,1212 @@ +(* Title: HOL/Rings.thy + Author: Gertrud Bauer + Author: Steven Obua + Author: Tobias Nipkow + Author: Lawrence C Paulson + Author: Markus Wenzel + Author: Jeremy Avigad +*) + +header {* Rings *} + +theory Rings +imports Groups +begin + +text {* + The theory of partially ordered rings is taken from the books: + \begin{itemize} + \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 + \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963 + \end{itemize} + Most of the used notions can also be looked up in + \begin{itemize} + \item \url{http://www.mathworld.com} by Eric Weisstein et. al. + \item \emph{Algebra I} by van der Waerden, Springer. + \end{itemize} +*} + +class semiring = ab_semigroup_add + semigroup_mult + + assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c" + assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c" +begin + +text{*For the @{text combine_numerals} simproc*} +lemma combine_common_factor: + "a * e + (b * e + c) = (a + b) * e + c" +by (simp add: left_distrib add_ac) + +end + +class mult_zero = times + zero + + assumes mult_zero_left [simp]: "0 * a = 0" + assumes mult_zero_right [simp]: "a * 0 = 0" + +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: left_distrib [symmetric]) + thus "0 * a = 0" by (simp only: add_left_cancel) +next + fix a :: 'a + have "a * 0 + a * 0 = a * 0 + 0" by (simp add: right_distrib [symmetric]) + thus "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: mult_ac) + also have "... = b * a + c * a" by (simp only: distrib) + also have "... = a * b + a * c" by (simp add: mult_ac) + 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) + +end + +class semiring_1 = zero_neq_one + semiring_0 + monoid_mult + +text {* Abstract divisibility *} + +class dvd = times +begin + +definition dvd :: "'a \ 'a \ bool" (infixl "dvd" 50) where + [code del]: "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 + +class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult + dvd + (*previously almost_semiring*) +begin + +subclass semiring_1 .. + +lemma dvd_refl[simp]: "a dvd a" +proof + show "a = a * 1" by simp +qed + +lemma dvd_trans: + assumes "a dvd b" and "b dvd c" + shows "a dvd c" +proof - + from assms obtain v where "b = a * v" by (auto elim!: dvdE) + moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE) + ultimately have "c = a * (v * w)" by (simp add: mult_assoc) + then show ?thesis .. +qed + +lemma dvd_0_left_iff [noatp, simp]: "0 dvd a \ a = 0" +by (auto intro: dvd_refl elim!: dvdE) + +lemma dvd_0_right [iff]: "a dvd 0" +proof + show "0 = a * 0" by simp +qed + +lemma one_dvd [simp]: "1 dvd a" +by (auto intro!: dvdI) + +lemma dvd_mult[simp]: "a dvd c \ a dvd (b * c)" +by (auto intro!: mult_left_commute dvdI elim!: dvdE) + +lemma dvd_mult2[simp]: "a dvd b \ a dvd (b * c)" + apply (subst mult_commute) + apply (erule dvd_mult) + done + +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: mult_ac) + 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" + unfolding mult_ac [of a] by (rule dvd_mult_left) + +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: right_distrib) + then show ?thesis .. +qed + +end + + +class no_zero_divisors = zero + times + + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" + +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 +begin + +subclass semiring_1_cancel .. +subclass comm_semiring_0_cancel .. +subclass comm_semiring_1 .. + +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: left_distrib [symmetric]) + +lemma minus_mult_right: "- (a * b) = a * - b" +by (rule minus_unique) (simp add: right_distrib [symmetric]) + +text{*Extract signs from products*} +lemmas mult_minus_left [simp, noatp] = minus_mult_left [symmetric] +lemmas mult_minus_right [simp,noatp] = 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]: "a * (b - c) = a * b - a * c" +by (simp add: right_distrib diff_minus) + +lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c" +by (simp add: left_distrib diff_minus) + +lemmas ring_distribs[noatp] = + right_distrib left_distrib left_diff_distrib right_diff_distrib + +text{*Legacy - use @{text algebra_simps} *} +lemmas ring_simps[noatp] = algebra_simps + +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[noatp] = + right_distrib left_distrib left_diff_distrib right_diff_distrib + +class comm_ring = comm_semiring + ab_group_add +begin + +subclass ring .. +subclass comm_semiring_0_cancel .. + +end + +class ring_1 = ring + zero_neq_one + monoid_mult +begin + +subclass semiring_1_cancel .. + +end + +class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult + (*previously ring*) +begin + +subclass ring_1 .. +subclass comm_semiring_1_cancel .. + +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)" +by (simp add: diff_minus dvd_minus_iff) + +end + +class ring_no_zero_divisors = ring + no_zero_divisors +begin + +lemma mult_eq_0_iff [simp]: + shows "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 + +text{*Cancellation of equalities with a common factor*} +lemma mult_cancel_right [simp, noatp]: + "a * c = b * c \ c = 0 \ a = b" +proof - + have "(a * c = b * c) = ((a - b) * c = 0)" + by (simp add: algebra_simps right_minus_eq) + thus ?thesis by (simp add: disj_commute right_minus_eq) +qed + +lemma mult_cancel_left [simp, noatp]: + "c * a = c * b \ c = 0 \ a = b" +proof - + have "(c * a = c * b) = (c * (a - b) = 0)" + by (simp add: algebra_simps right_minus_eq) + thus ?thesis by (simp add: right_minus_eq) +qed + +end + +class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors +begin + +lemma mult_cancel_right1 [simp]: + "c = b * c \ c = 0 \ b = 1" +by (insert mult_cancel_right [of 1 c b], force) + +lemma mult_cancel_right2 [simp]: + "a * c = c \ c = 0 \ a = 1" +by (insert mult_cancel_right [of a c 1], simp) + +lemma mult_cancel_left1 [simp]: + "c = c * b \ c = 0 \ b = 1" +by (insert mult_cancel_left [of c 1 b], force) + +lemma mult_cancel_left2 [simp]: + "c * a = c \ c = 0 \ a = 1" +by (insert mult_cancel_left [of c a 1], simp) + +end + +class idom = comm_ring_1 + no_zero_divisors +begin + +subclass ring_1_no_zero_divisors .. + +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: right_minus_eq eq_neg_iff_add_eq_0) +next + assume "a = b \ a = - b" + then show "a * a = b * b" by auto +qed + +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)" + unfolding dvd_def by (simp add: mult_ac) + also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" + unfolding dvd_def by simp + finally show ?thesis . +qed + +lemma dvd_mult_cancel_left [simp]: + "c * a dvd c * b \ c = 0 \ a dvd b" +proof - + have "c * a dvd c * b \ (\k. b * c = (a * k) * c)" + unfolding dvd_def by (simp add: mult_ac) + also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" + unfolding dvd_def by simp + finally show ?thesis . +qed + +end + +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" +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: + "inverse a = 0 \ a = 0" +apply (rule classical) +apply (drule nonzero_imp_inverse_nonzero) +apply auto +done + +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) + +end + +class mult_mono = times + zero + ord + + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" + assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" + +class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add +begin + +lemma mult_mono: + "a \ b \ c \ d \ 0 \ b \ 0 \ c + \ a * c \ b * d" +apply (erule mult_right_mono [THEN order_trans], assumption) +apply (erule mult_left_mono, assumption) +done + +lemma mult_mono': + "a \ b \ c \ d \ 0 \ a \ 0 \ c + \ a * c \ b * d" +apply (rule mult_mono) +apply (fast intro: order_trans)+ +done + +end + +class ordered_cancel_semiring = mult_mono + ordered_ab_semigroup_add + + semiring + cancel_comm_monoid_add +begin + +subclass semiring_0_cancel .. +subclass ordered_semiring .. + +lemma mult_nonneg_nonneg: "0 \ a \ 0 \ b \ 0 \ a * b" +using mult_left_mono [of zero b a] by simp + +lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" +using mult_left_mono [of b zero a] by simp + +lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" +using mult_right_mono [of a zero b] by simp + +text {* Legacy - use @{text mult_nonpos_nonneg} *} +lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" +by (drule mult_right_mono [of b zero], 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 linordered_semiring = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + mult_mono +begin + +subclass ordered_cancel_semiring .. + +subclass ordered_comm_monoid_add .. + +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 linordered_semiring_1 = linordered_semiring + semiring_1 + +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: "a \ b" "0 \ c" + from A show "c * a \ c * b" + unfolding le_less + using mult_strict_left_mono by (cases "c = 0") auto + from A 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 (force simp add: mult_strict_left_mono _not_less [symmetric]) + +lemma mult_right_le_imp_le: + "a * c \ b * c \ 0 < c \ a \ b" +by (force simp add: mult_strict_right_mono not_less [symmetric]) + +lemma mult_pos_pos: "0 < a \ 0 < b \ 0 < a * b" +using mult_strict_left_mono [of zero b a] by simp + +lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" +using mult_strict_left_mono [of b zero a] by simp + +lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" +using mult_strict_right_mono [of a zero b] by simp + +text {* Legacy - use @{text mult_neg_pos} *} +lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" +by (drule mult_strict_right_mono [of b zero], auto) + +lemma zero_less_mult_pos: + "0 < a * b \ 0 < a \ 0 < b" +apply (cases "b\0") + apply (auto simp add: le_less not_less) +apply (drule_tac mult_pos_neg [of a b]) + apply (auto dest: less_not_sym) +done + +lemma zero_less_mult_pos2: + "0 < b * a \ 0 < a \ 0 < b" +apply (cases "b\0") + apply (auto simp add: le_less not_less) +apply (drule_tac mult_pos_neg2 [of a b]) + apply (auto dest: less_not_sym) +done + +text{*Strict monotonicity in both arguments*} +lemma mult_strict_mono: + assumes "a < b" and "c < d" and "0 < b" and "0 \ c" + shows "a * c < b * d" + using assms apply (cases "c=0") + apply (simp add: mult_pos_pos) + apply (erule mult_strict_right_mono [THEN less_trans]) + apply (force simp add: le_less) + apply (erule mult_strict_left_mono, assumption) + done + +text{*This weaker variant has more natural premises*} +lemma mult_strict_mono': + assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" + shows "a * c < b * d" +by (rule mult_strict_mono) (insert assms, auto) + +lemma mult_less_le_imp_less: + assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" + shows "a * c < b * d" + using assms apply (subgoal_tac "a * c < b * c") + apply (erule less_le_trans) + apply (erule mult_left_mono) + apply simp + apply (erule mult_strict_right_mono) + apply assumption + done + +lemma mult_le_less_imp_less: + assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" + shows "a * c < b * d" + using assms apply (subgoal_tac "a * c \ b * c") + apply (erule le_less_trans) + apply (erule mult_strict_left_mono) + apply simp + apply (erule mult_right_mono) + apply simp + done + +lemma mult_less_imp_less_left: + assumes less: "c * a < c * b" and nonneg: "0 \ c" + shows "a < b" +proof (rule ccontr) + assume "\ a < b" + hence "b \ a" by (simp add: linorder_not_less) + hence "c * b \ c * a" using nonneg by (rule mult_left_mono) + with this and less show False by (simp add: not_less [symmetric]) +qed + +lemma mult_less_imp_less_right: + assumes less: "a * c < b * c" and nonneg: "0 \ c" + shows "a < b" +proof (rule ccontr) + assume "\ a < b" + hence "b \ a" by (simp add: linorder_not_less) + hence "b * c \ a * c" using nonneg by (rule mult_right_mono) + with this and less show False by (simp add: not_less [symmetric]) +qed + +end + +class linlinordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + +class mult_mono1 = times + zero + ord + + assumes mult_mono1: "a \ b \ 0 \ c \ c * a \ c * b" + +class ordered_comm_semiring = comm_semiring_0 + + ordered_ab_semigroup_add + mult_mono1 +begin + +subclass ordered_semiring +proof + fix a b c :: 'a + assume "a \ b" "0 \ c" + thus "c * a \ c * b" by (rule mult_mono1) + thus "a * c \ b * c" by (simp only: mult_commute) +qed + +end + +class ordered_cancel_comm_semiring = comm_semiring_0_cancel + + ordered_ab_semigroup_add + mult_mono1 +begin + +subclass ordered_comm_semiring .. +subclass ordered_cancel_semiring .. + +end + +class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + + assumes mult_strict_left_mono_comm: "a < b \ 0 < c \ c * a < c * b" +begin + +subclass linordered_semiring_strict +proof + fix a b c :: 'a + assume "a < b" "0 < c" + thus "c * a < c * b" by (rule mult_strict_left_mono_comm) + thus "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" + thus "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 .. + +text{*Legacy - use @{text algebra_simps} *} +lemmas ring_simps[noatp] = algebra_simps + +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" + apply (drule mult_left_mono [of _ _ "uminus c"]) + apply (simp_all add: minus_mult_left [symmetric]) + done + +lemma mult_right_mono_neg: + "b \ a \ c \ 0 \ a * c \ b * c" + apply (drule mult_right_mono [of _ _ "uminus c"]) + apply (simp_all add: minus_mult_right [symmetric]) + done + +lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" +using mult_right_mono_neg [of a zero b] by simp + +lemma split_mult_pos_le: + "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" +by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos) + +end + +class abs_if = minus + uminus + ord + zero + abs + + assumes abs_if: "\a\ = (if a < 0 then - a else a)" + +class sgn_if = minus + uminus + zero + one + ord + sgn + + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" + +lemma (in sgn_if) sgn0[simp]: "sgn 0 = 0" +by(simp add:sgn_if) + +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_less neg_less_eq_nonneg less_eq_neg_nonpos) + (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric] + neg_less_eq_nonneg less_eq_neg_nonpos, auto intro: add_nonneg_nonneg, + auto intro!: less_imp_le add_neg_neg) +qed (auto simp add: abs_if less_eq_neg_nonpos neg_equal_zero) + +end + +(* The "strict" suffix can be seen as describing the combination of linordered_ring and no_zero_divisors. + Basically, linordered_ring + no_zero_divisors = linordered_ring_strict. + *) +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 zero 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 note A' = this + show ?thesis proof (cases "b < 0") + case True with A' + show ?thesis by (auto dest: mult_neg_neg) + next + case False with B have "0 < b" by auto + with A' show ?thesis by (auto dest: mult_strict_right_mono) + qed + next + case False with A have A': "0 < a" by auto + show ?thesis proof (cases "b < 0") + case True with A' + show ?thesis by (auto dest: mult_strict_right_mono_neg) + next + case False with B have "0 < b" by auto + with A' show ?thesis by (auto dest: mult_pos_pos) + qed + qed + then show "a * b \ 0" by (simp add: neq_iff) +qed + +lemma zero_less_mult_iff: + "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" + apply (auto simp add: mult_pos_pos mult_neg_neg) + apply (simp_all add: not_less le_less) + apply (erule disjE) apply assumption defer + apply (erule disjE) defer apply (drule sym) apply simp + apply (erule disjE) defer apply (drule sym) apply simp + apply (erule disjE) apply assumption apply (drule sym) apply simp + apply (drule sym) apply simp + apply (blast dest: zero_less_mult_pos) + apply (blast dest: zero_less_mult_pos2) + done + +lemma zero_le_mult_iff: + "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: + "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" + apply (insert zero_less_mult_iff [of "-a" b]) + apply (force simp add: minus_mult_left[symmetric]) + done + +lemma mult_le_0_iff: + "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" + apply (insert zero_le_mult_iff [of "-a" b]) + apply (force simp add: minus_mult_left[symmetric]) + done + +lemma zero_le_square [simp]: "0 \ a * a" +by (simp add: zero_le_mult_iff linear) + +lemma not_square_less_zero [simp]: "\ (a * a < 0)" +by (simp add: not_less) + +text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"}, + also with the relations @{text "\"} 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" + apply (cases "c = 0") + apply (auto simp add: neq_iff mult_strict_right_mono + mult_strict_right_mono_neg) + apply (auto simp add: not_less + not_le [symmetric, of "a*c"] + not_le [symmetric, of a]) + apply (erule_tac [!] notE) + apply (auto simp add: less_imp_le mult_right_mono + mult_right_mono_neg) + done + +lemma mult_less_cancel_left_disj: + "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" + apply (cases "c = 0") + apply (auto simp add: neq_iff mult_strict_left_mono + mult_strict_left_mono_neg) + apply (auto simp add: not_less + not_le [symmetric, of "c*a"] + not_le [symmetric, of a]) + apply (erule_tac [!] notE) + apply (auto simp add: less_imp_le mult_left_mono + mult_left_mono_neg) + done + +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 + +text{*Legacy - use @{text algebra_simps} *} +lemmas ring_simps[noatp] = algebra_simps + +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_semidom = comm_semiring_1_cancel + linordered_comm_semiring_strict + + (*previously linordered_semiring*) + assumes zero_less_one [simp]: "0 < 1" +begin + +lemma pos_add_strict: + shows "0 < a \ b < c \ b < a + c" + using add_strict_mono [of zero a b c] by simp + +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 less_1_mult: + assumes "1 < m" and "1 < n" + shows "1 < m * n" + using assms 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_if + (*previously linordered_ring*) +begin + +subclass linordered_ring_strict .. +subclass ordered_comm_ring .. +subclass idom .. + +subclass linordered_semidom +proof + have "0 \ 1 * 1" by (rule zero_le_square) + thus "0 < 1" by (simp add: le_less) +qed + +lemma linorder_neqE_linordered_idom: + assumes "x \ y" obtains "x < y" | "y < x" + using assms by (rule neqE) + +text {* These cancellation simprules 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)" +by (insert mult_le_cancel_right [of 1 c b], simp) + +lemma mult_le_cancel_right2: + "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" +by (insert mult_le_cancel_right [of a c 1], simp) + +lemma mult_le_cancel_left1: + "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" +by (insert mult_le_cancel_left [of c 1 b], simp) + +lemma mult_le_cancel_left2: + "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" +by (insert mult_le_cancel_left [of c a 1], simp) + +lemma mult_less_cancel_right1: + "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" +by (insert mult_less_cancel_right [of 1 c b], simp) + +lemma mult_less_cancel_right2: + "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" +by (insert mult_less_cancel_right [of a c 1], simp) + +lemma mult_less_cancel_left1: + "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" +by (insert mult_less_cancel_left [of c 1 b], simp) + +lemma mult_less_cancel_left2: + "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" +by (insert mult_less_cancel_left [of c a 1], simp) + +lemma sgn_sgn [simp]: + "sgn (sgn a) = sgn a" +unfolding sgn_if by simp + +lemma sgn_0_0: + "sgn a = 0 \ a = 0" +unfolding sgn_if by simp + +lemma sgn_1_pos: + "sgn a = 1 \ a > 0" +unfolding sgn_if by (simp add: neg_equal_zero) + +lemma sgn_1_neg: + "sgn a = - 1 \ a < 0" +unfolding sgn_if by (auto simp add: equal_neg_zero) + +lemma sgn_pos [simp]: + "0 < a \ sgn a = 1" +unfolding sgn_1_pos . + +lemma sgn_neg [simp]: + "a < 0 \ sgn a = - 1" +unfolding sgn_1_neg . + +lemma sgn_times: + "sgn (a * b) = sgn a * sgn b" +by (auto simp add: sgn_if zero_less_mult_iff) + +lemma abs_sgn: "abs 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_dvd_iff [simp]: "(abs m) dvd k \ m dvd k" + by (simp add: abs_if) + +lemma dvd_abs_iff [simp]: "m dvd (abs k) \ m dvd k" + by (simp add: abs_if) + +lemma dvd_if_abs_eq: + "abs l = abs (k) \ l dvd k" +by(subst abs_dvd_iff[symmetric]) simp + +end + +text {* Simprules for comparisons where common factors can be cancelled. *} + +lemmas mult_compare_simps[noatp] = + mult_le_cancel_right mult_le_cancel_left + mult_le_cancel_right1 mult_le_cancel_right2 + mult_le_cancel_left1 mult_le_cancel_left2 + mult_less_cancel_right mult_less_cancel_left + mult_less_cancel_right1 mult_less_cancel_right2 + mult_less_cancel_left1 mult_less_cancel_left2 + mult_cancel_right mult_cancel_left + mult_cancel_right1 mult_cancel_right2 + mult_cancel_left1 mult_cancel_left2 + +-- {* FIXME continue localization here *} + +subsection {* Reasoning about inequalities with division *} + +lemma mult_right_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1 + ==> x * y <= x" +by (auto simp add: mult_compare_simps) + +lemma mult_left_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1 + ==> y * x <= x" +by (auto simp add: mult_compare_simps) + +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) + thus ?thesis by simp +qed + +lemma zero_less_two: "0 < 1 + 1" +by (blast intro: less_trans zero_less_one less_add_one) + +end + + +subsection {* Absolute Value *} + +context linordered_idom +begin + +lemma mult_sgn_abs: "sgn x * abs x = x" + unfolding abs_if sgn_if by auto + +end + +lemma abs_one [simp]: "abs 1 = (1::'a::linordered_idom)" +by (simp add: abs_if zero_less_one [THEN order_less_not_sym]) + +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 proof +qed (auto simp add: abs_if not_less equal_neg_zero neg_equal_zero mult_less_0_iff) + +lemma abs_mult: + "abs (a * b) = abs a * abs b" + by (rule abs_eq_mult) auto + +lemma abs_mult_self: + "abs a * abs a = a * a" + by (simp add: abs_if) + +end + +lemma abs_mult_less: + "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::linordered_idom)" +proof - + assume ac: "abs a < c" + hence cpos: "0 + (abs y) * x = abs (y * x)" + apply (subst abs_mult) + apply simp +done + +code_modulename SML + Rings Arith + +code_modulename OCaml + Rings Arith + +code_modulename Haskell + Rings Arith + +end diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Tools/Qelim/cooper.ML Mon Feb 08 17:12:38 2010 +0100 @@ -79,9 +79,9 @@ | Const (@{const_name Algebras.less_eq}, _) $ y $ z => if term_of x aconv y then Le (Thm.dest_arg ct) else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox -| Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_) => +| Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_) => if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox -| Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_)) => +| Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_)) => if term_of x aconv y then NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox | _ => Nox) @@ -222,8 +222,8 @@ | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name Algebras.less_eq}, T) $ s $ t)) = lin vs (Const (@{const_name Algebras.less}, T) $ t $ s) | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t) - | lin (vs as x::_) (Const(@{const_name Ring_and_Field.dvd},_)$d$t) = - HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} (numeral1 abs d, lint vs t) + | lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) = + HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t) | lin (vs as x::_) ((b as Const("op =",_))$s$t) = (case lint vs (subC$t$s) of (t as a$(m$c$y)$r) => @@ -253,7 +253,7 @@ | is_intrel _ = false; fun linearize_conv ctxt vs ct = case term_of ct of - Const(@{const_name Ring_and_Field.dvd},_)$d$t => + Const(@{const_name Rings.dvd},_)$d$t => let val th = binop_conv (lint_conv ctxt vs) ct val (d',t') = Thm.dest_binop (Thm.rhs_of th) @@ -278,7 +278,7 @@ | _ => dth end end -| Const (@{const_name Not},_)$(Const(@{const_name Ring_and_Field.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct +| Const (@{const_name Not},_)$(Const(@{const_name Rings.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct | t => if is_intrel t then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop)) RS eq_reflection @@ -301,7 +301,7 @@ if x aconv y andalso member (op =) [@{const_name Algebras.less}, @{const_name Algebras.less_eq}] s then (ins (dest_numeral c) acc, dacc) else (acc,dacc) - | Const(@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_) => + | Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_) => if x aconv y then (acc,ins (dest_numeral c) dacc) else (acc,dacc) | Const("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t) | Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t) @@ -343,7 +343,7 @@ if x=y andalso member (op =) [@{const_name Algebras.less}, @{const_name Algebras.less_eq}] s then cv (l div dest_numeral c) t else Thm.reflexive t - | Const(@{const_name Ring_and_Field.dvd},_)$d$(r as (Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_)) => + | Const(@{const_name Rings.dvd},_)$d$(r as (Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_)) => if x=y then let val k = l div dest_numeral c @@ -562,7 +562,7 @@ | Const("False",_) => F | Const(@{const_name Algebras.less},_)$t1$t2 => Lt (Sub (i_of_term vs t1,i_of_term vs t2)) | Const(@{const_name Algebras.less_eq},_)$t1$t2 => Le (Sub(i_of_term vs t1,i_of_term vs t2)) - | Const(@{const_name Ring_and_Field.dvd},_)$t1$t2 => + | Const(@{const_name Rings.dvd},_)$t1$t2 => (Dvd(HOLogic.dest_number t1 |> snd, i_of_term vs t2) handle _ => cooper "Reification: unsupported dvd") (* FIXME avoid handle _ *) | @{term "op = :: int => _"}$t1$t2 => Eq (Sub (i_of_term vs t1,i_of_term vs t2)) | @{term "op = :: bool => _ "}$t1$t2 => Iff(qf_of_term ps vs t1,qf_of_term ps vs t2) diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Tools/lin_arith.ML Mon Feb 08 17:12:38 2010 +0100 @@ -236,7 +236,7 @@ end handle Rat.DIVZERO => NONE; fun of_lin_arith_sort thy U = - Sign.of_sort thy (U, @{sort Ring_and_Field.linordered_idom}); + Sign.of_sort thy (U, @{sort Rings.linordered_idom}); fun allows_lin_arith thy (discrete : string list) (U as Type (D, [])) : bool * bool = if of_lin_arith_sort thy U then (true, member (op =) discrete D) @@ -814,8 +814,8 @@ addsimps @{thms ring_distribs} addsimps [@{thm if_True}, @{thm if_False}] addsimps - [@{thm "monoid_add_class.add_0_left"}, - @{thm "monoid_add_class.add_0_right"}, + [@{thm add_0_left}, + @{thm add_0_right}, @{thm "Zero_not_Suc"}, @{thm "Suc_not_Zero"}, @{thm "le_0_eq"}, @{thm "One_nat_def"}, @{thm "order_less_irrefl"}, @{thm "zero_neq_one"}, @{thm "zero_less_one"}, @{thm "zero_le_one"}, @{thm "zero_neq_one"} RS not_sym, @{thm "not_one_le_zero"}, diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Tools/nat_numeral_simprocs.ML --- a/src/HOL/Tools/nat_numeral_simprocs.ML Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Tools/nat_numeral_simprocs.ML Mon Feb 08 17:12:38 2010 +0100 @@ -124,7 +124,7 @@ (*Simplify 1*n and n*1 to n*) -val add_0s = map rename_numerals [@{thm add_0}, @{thm add_0_right}]; +val add_0s = map rename_numerals [@{thm add_0}, @{thm Nat.add_0_right}]; val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}]; (*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*) @@ -136,7 +136,7 @@ val simplify_meta_eq = Arith_Data.simplify_meta_eq - ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right}, + ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm Nat.add_0_right}, @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules); @@ -290,8 +290,8 @@ structure DvdCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val prove_conv = Arith_Data.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT + val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT val cancel = @{thm nat_mult_dvd_cancel1} RS trans val neg_exchanges = false ) @@ -411,8 +411,8 @@ structure DvdCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val prove_conv = Arith_Data.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT + val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT fun simp_conv _ _ = SOME @{thm nat_mult_dvd_cancel_disj} ); diff -r 00f311c32444 -r 9f841f20dca6 src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Mon Feb 08 17:12:32 2010 +0100 +++ b/src/HOL/Tools/numeral_simprocs.ML Mon Feb 08 17:12:38 2010 +0100 @@ -373,7 +373,7 @@ [@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}] @ simps fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) val simplify_meta_eq = Arith_Data.simplify_meta_eq - [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left}, + [@{thm add_0}, @{thm Nat.add_0_right}, @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}]; end @@ -561,8 +561,8 @@ structure DvdCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val prove_conv = Arith_Data.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} Term.dummyT + val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} Term.dummyT fun simp_conv _ _ = SOME @{thm dvd_mult_cancel_left} );