# HG changeset patch # User haftmann # Date 1193730354 -3600 # Node ID 022029099a83e5eaa562e7eabe173abb771617a5 # Parent 2673709fb8f770a9791909c1c5371ee984c6a575 continued localization diff -r 2673709fb8f7 -r 022029099a83 src/HOL/Hyperreal/StarClasses.thy --- a/src/HOL/Hyperreal/StarClasses.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/Hyperreal/StarClasses.thy Tue Oct 30 08:45:54 2007 +0100 @@ -403,7 +403,7 @@ done instance star :: (pordered_comm_semiring) pordered_comm_semiring -by (intro_classes, transfer, rule mult_mono1_class.times_zero_less_eq_less.mult_mono) +by (intro_classes, transfer, rule mult_mono1_class.times_zero_less_eq_less.mult_mono1) instance star :: (pordered_cancel_comm_semiring) pordered_cancel_comm_semiring .. @@ -451,7 +451,7 @@ subsection {* Number classes *} lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" -by (induct_tac n, simp_all) +by (induct n, simp_all) lemma Standard_of_nat [simp]: "of_nat n \ Standard" by (simp add: star_of_nat_def) diff -r 2673709fb8f7 -r 022029099a83 src/HOL/IntArith.thy --- a/src/HOL/IntArith.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/IntArith.thy Tue Oct 30 08:45:54 2007 +0100 @@ -208,9 +208,20 @@ with False show ?thesis by simp qed -lemma of_int_of_nat: "of_int k = (if k < 0 then - of_nat (nat (- k)) - else of_nat (nat k))" - by (auto simp add: of_nat_nat) +context ring_1 +begin + +lemma of_int_of_nat: + "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" +proof (cases "k < 0") + case True then have "0 \ - k" by simp + then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) + with True show ?thesis by simp +next + case False then show ?thesis by (simp add: not_less of_nat_nat) +qed + +end lemma nat_mult_distrib: "(0::int) \ z ==> nat (z*z') = nat z * nat z'" apply (cases "0 \ z'") diff -r 2673709fb8f7 -r 022029099a83 src/HOL/IntDef.thy --- a/src/HOL/IntDef.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/IntDef.thy Tue Oct 30 08:45:54 2007 +0100 @@ -429,6 +429,8 @@ context ring_1 begin +term of_nat + definition of_int :: "int \ 'a" where @@ -518,10 +520,15 @@ by (cases z) (simp add: of_int add minus int_def diff_minus) qed +context ring_1 +begin + lemma of_nat_nat: "0 \ z \ of_nat (nat z) = of_int z" -by (cases z rule: eq_Abs_Integ) + by (cases z rule: eq_Abs_Integ) (simp add: nat le of_int Zero_int_def of_nat_diff) +end + subsection{*The Set of Integers*} diff -r 2673709fb8f7 -r 022029099a83 src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/OrderedGroup.thy Tue Oct 30 08:45:54 2007 +0100 @@ -535,8 +535,21 @@ lemma le_iff_diff_le_0: "a \ b \ a - b \ 0" by (simp add: compare_rls) +lemmas group_simps = + add_ac + add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2 + diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff + diff_less_eq less_diff_eq diff_le_eq le_diff_eq + end +lemmas group_simps = + mult_ac + add_ac + add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2 + diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff + diff_less_eq less_diff_eq diff_le_eq le_diff_eq + class ordered_ab_semigroup_add = linorder + pordered_ab_semigroup_add @@ -565,6 +578,12 @@ qed qed +class ordered_ab_group_add = + linorder + pordered_ab_group_add + +subclass (in ordered_ab_group_add) ordered_cancel_ab_semigroup_add + by unfold_locales + -- {* FIXME localize the following *} lemma add_increasing: @@ -752,6 +771,12 @@ (simp add: le_infI) qed +lemma neg_inf_eq_sup: "- inf a b = sup (-a) (-b)" + by (simp add: inf_eq_neg_sup) + +lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)" + by (simp add: sup_eq_neg_inf) + lemma add_eq_inf_sup: "a + b = sup a b + inf a b" proof - have "0 = - inf 0 (a-b) + inf (a-b) 0" by (simp add: inf_commute) @@ -776,6 +801,21 @@ pprt :: "'a \ 'a" where "pprt x = sup x 0" +lemma pprt_neg: "pprt (- x) = - nprt x" +proof - + have "sup (- x) 0 = sup (- x) (- 0)" unfolding minus_zero .. + also have "\ = - inf x 0" unfolding neg_inf_eq_sup .. + finally have "sup (- x) 0 = - inf x 0" . + then show ?thesis unfolding pprt_def nprt_def . +qed + +lemma nprt_neg: "nprt (- x) = - pprt x" +proof - + from pprt_neg have "pprt (- (- x)) = - nprt (- x)" . + then have "pprt x = - nprt (- x)" by simp + then show ?thesis by simp +qed + lemma prts: "a = pprt a + nprt a" by (simp add: pprt_def nprt_def add_eq_inf_sup[symmetric]) @@ -904,38 +944,67 @@ ultimately show ?thesis by blast qed +declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] + +lemma le_minus_self_iff: "a \ - a \ a \ 0" +proof - + from add_le_cancel_left [of "uminus a" "plus a a" zero] + 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 "uminus a" zero "plus a a"] + have "(-a <= a) = (0 <= a+a)" + by (simp add: add_assoc[symmetric]) + thus ?thesis by simp +qed + +lemma zero_le_iff_zero_nprt: "0 \ a \ nprt a = 0" + by (simp add: le_iff_inf nprt_def inf_commute) + +lemma le_zero_iff_zero_pprt: "a \ 0 \ pprt a = 0" + by (simp add: le_iff_sup pprt_def sup_commute) + +lemma le_zero_iff_pprt_id: "0 \ a \ pprt a = a" + by (simp add: le_iff_sup pprt_def sup_commute) + +lemma zero_le_iff_nprt_id: "a \ 0 \ nprt a = a" + by (simp add: le_iff_inf nprt_def inf_commute) + +lemma pprt_mono [simp, noatp]: "a \ b \ pprt a \ pprt b" + by (simp add: le_iff_sup pprt_def sup_ACI sup_assoc [symmetric, of a]) + +lemma nprt_mono [simp, noatp]: "a \ b \ nprt a \ nprt b" + by (simp add: le_iff_inf nprt_def inf_ACI inf_assoc [symmetric, of a]) + end lemmas add_sup_inf_distribs = add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left -class lordered_ab_group_abs = lordered_ab_group + abs + - assumes abs_lattice: "\x\ = sup x (- x)" + +class pordered_ab_group_add_abs = pordered_ab_group_add + abs + + assumes abs_ge_zero [simp]: "\a\ \ 0" + and abs_ge_self: "a \ \a\" + and abs_of_nonneg [simp]: "0 \ a \ \a\ = a" + and abs_leI: "a \ b \ - a \ b \ \a\ \ b" + and abs_eq_0 [simp]: "\a\ = 0 \ a = 0" + and abs_minus_cancel [simp]: "\-a\ = \a\" + and abs_idempotent [simp]: "\\a\\ = \a\" + and abs_triangle_ineq: "\a + b\ \ \a\ + \b\" begin lemma abs_zero [simp]: "\0\ = 0" - by (simp add: abs_lattice) - -lemma abs_eq_0 [simp]: "\a\ = 0 \ a = 0" - by (simp add: abs_lattice) + by simp lemma abs_0_eq [simp, noatp]: "0 = \a\ \ a = 0" proof - - have "(0 = abs a) = (abs a = 0)" by (simp only: eq_ac) + have "0 = \a\ \ \a\ = 0" by (simp only: eq_ac) thus ?thesis by simp qed -lemma neg_inf_eq_sup [simp]: "- inf a b = sup (-a) (-b)" - by (simp add: inf_eq_neg_sup) - -lemma neg_sup_eq_inf [simp]: "- sup a b = inf (-a) (-b)" - by (simp del: neg_inf_eq_sup add: sup_eq_neg_inf) - -lemma abs_ge_zero [simp]: "0 \ \a\" -proof - - have a: "a \ \a\" and b: "- a \ \a\" by (auto simp add: abs_lattice) - show ?thesis by (rule add_mono [OF a b, simplified]) -qed - lemma abs_le_zero_iff [simp]: "\a\ \ 0 \ a = 0" proof assume "\a\ \ 0" @@ -955,20 +1024,11 @@ show ?thesis by (simp add: a) qed -lemma abs_ge_self: "a \ \a\" - by (auto simp add: abs_lattice) - lemma abs_ge_minus_self: "- a \ \a\" - by (auto simp add: abs_lattice) - -lemma abs_minus_cancel [simp]: "\-a\ = \a\" - by (simp add: abs_lattice sup_commute) - -lemma abs_idempotent [simp]: "\\a\\ = \a\" -apply (simp add: abs_lattice [of "abs a"]) -apply (subst sup_absorb1) -apply (rule order_trans [of _ zero]) -by auto +proof - + have "- a \ \-a\" by (rule abs_ge_self) + then show ?thesis by simp +qed lemma abs_minus_commute: "\a - b\ = \b - a\" @@ -978,73 +1038,24 @@ finally show ?thesis . qed -lemma abs_prts: "\a\ = pprt a - nprt a" +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 - - have "0 \ \a\" by simp - then have "0 \ sup a (- a)" unfolding abs_lattice . - then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1) - then show ?thesis - by (simp add: add_sup_inf_distribs sup_ACI - pprt_def nprt_def diff_minus abs_lattice) + 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 zero_le_iff_zero_nprt: "0 \ a \ nprt a = 0" - by (simp add: le_iff_inf nprt_def inf_commute) - -lemma le_zero_iff_zero_pprt: "a \ 0 \ pprt a = 0" - by (simp add: le_iff_sup pprt_def sup_commute) - -lemma le_zero_iff_pprt_id: "0 \ a \ pprt a = a" - by (simp add: le_iff_sup pprt_def sup_commute) - -lemma zero_le_iff_nprt_id: "a \ 0 \ nprt a = a" -by (simp add: le_iff_inf nprt_def inf_commute) - -lemma pprt_mono [simp, noatp]: "a \ b \ pprt a \ pprt b" - by (simp add: le_iff_sup pprt_def sup_ACI sup_assoc [symmetric, of a]) - -lemma nprt_mono [simp, noatp]: "a \ b \ nprt a \ nprt b" - by (simp add: le_iff_inf nprt_def inf_ACI inf_assoc [symmetric, of a]) - -lemma pprt_neg: "pprt (- x) = - nprt x" - by (simp add: pprt_def nprt_def) - -lemma nprt_neg: "nprt (- x) = - pprt x" - by (simp add: pprt_def nprt_def) - -lemma abs_of_nonneg [simp]: "0 \ a \ \a\ = a" - by (simp add: iffD1 [OF zero_le_iff_zero_nprt] - iffD1[OF le_zero_iff_pprt_id] abs_prts) - -lemma abs_of_pos: "0 < x \ \x\ = x" - by (rule abs_of_nonneg, rule less_imp_le) - -lemma abs_of_nonpos [simp]: "a \ 0 \ \a\ = - a" - by (simp add: iffD1 [OF le_zero_iff_zero_pprt] - iffD1 [OF zero_le_iff_nprt_id] abs_prts) - -lemma abs_of_neg: "x < 0 \ \x\ = - x" +lemma abs_of_neg: "a < 0 \ \a\ = - a" by (rule abs_of_nonpos, rule less_imp_le) -lemma abs_leI: "a \ b \ - a \ b \ \a\ \ b" - by (simp add: abs_lattice le_supI) - -lemma le_minus_self_iff: "a \ - a \ a \ 0" -proof - - from add_le_cancel_left [of "uminus a" "plus a a" zero] - 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 "uminus a" zero "plus a a"] - have "(-a <= a) = (0 <= a+a)" - by (simp add: add_assoc[symmetric]) - thus ?thesis by simp -qed - lemma abs_le_D1: "\a\ \ b \ a \ b" by (insert abs_ge_self, blast intro: order_trans) @@ -1054,20 +1065,6 @@ 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_ineq: "\a + b\ \ \a\ + \b\" -proof - - have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n") - by (simp add: abs_lattice add_sup_inf_distribs sup_ACI diff_minus) - have a:"a+b <= sup ?m ?n" by (simp) - have b:"-a-b <= ?n" by (simp) - have c:"?n <= sup ?m ?n" by (simp) - from b c have d: "-a-b <= sup ?m ?n" by(rule order_trans) - have e:"-a-b = -(a+b)" by (simp add: diff_minus) - from a d e have "abs(a+b) <= sup ?m ?n" - by (drule_tac abs_leI, auto) - with g[symmetric] show ?thesis by simp -qed - lemma abs_triangle_ineq2: "\a\ - \b\ \ \a - b\" apply (simp add: compare_rls) apply (subgoal_tac "abs a = abs (plus (minus a b) b)") @@ -1102,7 +1099,7 @@ finally show ?thesis . qed -lemma abs_add_abs[simp]: +lemma abs_add_abs [simp]: "\\a\ + \b\\ = \a\ + \b\" (is "?L = ?R") proof (rule antisym) show "?L \ ?R" by(rule abs_ge_self) @@ -1114,6 +1111,87 @@ end + +class lordered_ab_group_abs = lordered_ab_group + abs + + assumes abs_lattice: "\a\ = sup a (- a)" +begin + +lemma abs_prts: "\a\ = pprt a - nprt a" +proof - + have "0 \ \a\" + proof - + have a: "a \ \a\" and b: "- a \ \a\" by (auto simp add: abs_lattice) + show ?thesis by (rule add_mono [OF a b, simplified]) + qed + then have "0 \ sup a (- a)" unfolding abs_lattice . + then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1) + then show ?thesis + by (simp add: add_sup_inf_distribs sup_ACI + pprt_def nprt_def diff_minus abs_lattice) +qed + +subclass pordered_ab_group_add_abs +proof - + have abs_ge_zero [simp]: "\a. 0 \ \a\" + proof - + fix a b + have a: "a \ \a\" and b: "- a \ \a\" by (auto simp add: abs_lattice) + show "0 \ \a\" by (rule add_mono [OF a b, simplified]) + qed + have abs_leI: "\a b. a \ b \ - a \ b \ \a\ \ b" + by (simp add: abs_lattice le_supI) + show ?thesis + proof unfold_locales + fix a + show "0 \ \a\" by simp + next + fix a + show "a \ \a\" + by (auto simp add: abs_lattice) + next + fix a + show "\a\ = 0 \ a = 0" + by (simp add: abs_lattice) + next + fix a + show "\-a\ = \a\" + by (simp add: abs_lattice sup_commute) + next + fix a + show "\\a\\ = \a\" + apply (simp add: abs_lattice [of "abs a"]) + apply (subst sup_absorb1) + apply (rule order_trans [of _ zero]) + apply auto + done + next + fix a + show "0 \ a \ \a\ = a" + by (simp add: iffD1 [OF zero_le_iff_zero_nprt] + iffD1[OF le_zero_iff_pprt_id] abs_prts) + next + fix a b + show "a \ b \ - a \ b \ \a\ \ b" by (erule abs_leI) + next + fix a b + show "\a + b\ \ \a\ + \b\" + proof - + have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n") + by (simp add: abs_lattice add_sup_inf_distribs sup_ACI diff_minus) + have a:"a+b <= sup ?m ?n" by (simp) + have b:"-a-b <= ?n" by (simp) + have c:"?n <= sup ?m ?n" by (simp) + from b c have d: "-a-b <= sup ?m ?n" by(rule order_trans) + have e:"-a-b = -(a+b)" by (simp add: diff_minus) + from a d e have "abs(a+b) <= sup ?m ?n" + by (drule_tac abs_leI, auto) + with g[symmetric] show ?thesis by simp + qed + qed auto +qed + +end + lemma sup_eq_if: fixes a :: "'a\{lordered_ab_group, linorder}" shows "sup a (- a) = (if a < 0 then - a else a)" @@ -1168,13 +1246,6 @@ apply (simp_all add: prems) done -lemmas group_simps = - mult_ac - add_ac - add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2 - diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff - diff_less_eq less_diff_eq diff_le_eq le_diff_eq - lemma estimate_by_abs: "a + b <= (c::'a::lordered_ab_group_abs) \ a <= c + abs b" proof - diff -r 2673709fb8f7 -r 022029099a83 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/Presburger.thy Tue Oct 30 08:45:54 2007 +0100 @@ -703,6 +703,8 @@ less_Bit0_Min less_Bit1_Min less_Bit_Bit0 less_Bit1_Bit less_Bit0_Bit1 less_number_of +context ring_1 +begin lemma of_int_num [code func]: "of_int k = (if k = 0 then 0 else if k < 0 then @@ -737,3 +739,5 @@ qed end + +end diff -r 2673709fb8f7 -r 022029099a83 src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Mon Oct 29 17:08:01 2007 +0100 +++ b/src/HOL/Ring_and_Field.thy Tue Oct 30 08:45:54 2007 +0100 @@ -153,17 +153,25 @@ lemmas ring_distribs = right_distrib left_distrib left_diff_distrib right_diff_distrib +lemmas ring_simps = + add_ac + add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2 + diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff + ring_distribs + +lemma eq_add_iff1: + "a * e + c = b * e + d \ (a - b) * e + c = d" + by (simp add: ring_simps) + +lemma eq_add_iff2: + "a * e + c = b * e + d \ c = (b - a) * e + d" + by (simp add: ring_simps) + end lemmas ring_distribs = right_distrib left_distrib left_diff_distrib right_diff_distrib -text{*This list of rewrites simplifies ring terms by multiplying -everything out and bringing sums and products into a canonical form -(by ordered rewriting). As a result it decides ring equalities but -also helps with inequalities. *} -lemmas ring_simps = group_simps ring_distribs - class comm_ring = comm_semiring + ab_group_add subclass (in comm_ring) ring by unfold_locales @@ -180,6 +188,18 @@ subclass (in comm_ring_1) comm_semiring_1_cancel by unfold_locales 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 + +end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors @@ -226,16 +246,75 @@ thus "inverse a * a = 1" by (rule field_inverse) thus "a * inverse a = 1" by (simp only: mult_commute) qed + subclass (in field) idom by unfold_locales +context field +begin + +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 ring_distribs) + +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 add: divide_self) + 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 pordered_semiring = mult_mono + semiring_0 + pordered_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 pordered_cancel_semiring = mult_mono + pordered_ab_semigroup_add + semiring + comm_monoid_add + cancel_ab_semigroup_add @@ -243,10 +322,37 @@ subclass (in pordered_cancel_semiring) semiring_0_cancel by unfold_locales subclass (in pordered_cancel_semiring) pordered_semiring by unfold_locales -class ordered_semiring = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add + mult_mono +context pordered_cancel_semiring begin -subclass pordered_cancel_semiring by unfold_locales +lemma mult_nonneg_nonneg: "0 \ a \ 0 \ b \ 0 \ a * b" + by (drule mult_left_mono [of zero b], auto) + +lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" + by (drule mult_left_mono [of b zero], auto) + +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::_::pordered_cancel_semiring)" + by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) + +end + +class ordered_semiring = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add + mult_mono + +subclass (in ordered_semiring) pordered_cancel_semiring by unfold_locales + +context ordered_semiring +begin + +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 @@ -268,8 +374,49 @@ using mult_strict_right_mono by (cases "c = 0") auto qed +context ordered_semiring_strict +begin + +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" + by (drule mult_strict_left_mono [of zero b], auto) + +lemma mult_pos_neg: + "0 < a \ b < 0 \ a * b < 0" + by (drule mult_strict_left_mono [of b zero], auto) + +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 + +end + class mult_mono1 = times + zero + ord + - assumes mult_mono: "a \ b \ 0 \ c \ c * a \ c * b" + assumes mult_mono1: "a \ b \ 0 \ c \ c * a \ c * b" class pordered_comm_semiring = comm_semiring_0 + pordered_ab_semigroup_add + mult_mono1 @@ -289,7 +436,7 @@ proof unfold_locales fix a b c :: 'a assume "a \ b" "0 \ c" - thus "c * a \ c * b" by (rule mult_mono) + thus "c * a \ c * b" by (rule mult_mono1) thus "a * c \ b * c" by (simp only: mult_commute) qed @@ -314,9 +461,49 @@ qed class pordered_ring = ring + pordered_cancel_semiring + +subclass (in pordered_ring) pordered_ab_group_add by unfold_locales + +context pordered_ring begin -subclass pordered_ab_group_add by unfold_locales +lemmas ring_simps = ring_simps group_simps + +lemma less_add_iff1: + "a * e + c < b * e + d \ (a - b) * e + c < d" + by (simp add: ring_simps) + +lemma less_add_iff2: + "a * e + c < b * e + d \ c < (b - a) * e + d" + by (simp add: ring_simps) + +lemma le_add_iff1: + "a * e + c \ b * e + d \ (a - b) * e + c \ d" + by (simp add: ring_simps) + +lemma le_add_iff2: + "a * e + c \ b * e + d \ c \ (b - a) * e + d" + by (simp add: ring_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" + by (drule mult_right_mono_neg [of a zero b]) auto + +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 @@ -331,12 +518,10 @@ class sgn_if = sgn + zero + one + minus + ord + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" -(* The "strict" suffix can be seen as describing the combination of ordered_ring and no_zero_divisors. - Basically, ordered_ring + no_zero_divisors = ordered_ring_strict. - *) -class ordered_ring = ring + ordered_semiring + lordered_ab_group + abs_if - --- {*FIXME: continue localization here*} +class ordered_ring = ring + ordered_semiring + + lordered_ab_group + abs_if + -- {*FIXME: should inherit from ordered_ab_group_add rather than + lordered_ab_group*} instance ordered_ring \ lordered_ring proof @@ -345,167 +530,44 @@ by (simp only: abs_if sup_eq_if) qed -class ordered_ring_strict = - ring + ordered_semiring_strict + lordered_ab_group + abs_if - -instance ordered_ring_strict \ ordered_ring .. - -class pordered_comm_ring = comm_ring + pordered_comm_semiring - -instance pordered_comm_ring \ pordered_ring .. - -instance pordered_comm_ring \ pordered_cancel_comm_semiring .. - -class ordered_semidom = comm_semiring_1_cancel + ordered_comm_semiring_strict + - (*previously ordered_semiring*) - assumes zero_less_one [simp]: "0 < 1" - -lemma pos_add_strict: - fixes a b c :: "'a\ordered_semidom" - shows "0 < a \ b < c \ b < a + c" - using add_strict_mono [of 0 a b c] by simp - -class ordered_idom = - comm_ring_1 + - ordered_comm_semiring_strict + - lordered_ab_group + - abs_if + sgn_if - (*previously ordered_ring*) - -instance ordered_idom \ ordered_ring_strict .. - -instance ordered_idom \ pordered_comm_ring .. - -class ordered_field = field + ordered_idom - -lemma linorder_neqE_ordered_idom: - fixes x y :: "'a :: ordered_idom" - assumes "x \ y" obtains "x < y" | "y < x" - using assms by (rule linorder_neqE) +(* The "strict" suffix can be seen as describing the combination of ordered_ring and no_zero_divisors. + Basically, ordered_ring + no_zero_divisors = ordered_ring_strict. + *) +class ordered_ring_strict = ring + ordered_semiring_strict + + lordered_ab_group + abs_if + -- {*FIXME: should inherit from ordered_ab_group_add rather than + lordered_ab_group*} -lemma eq_add_iff1: - "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))" -by (simp add: ring_simps) - -lemma eq_add_iff2: - "(a*e + c = b*e + d) = (c = (b-a)*e + (d::'a::ring))" -by (simp add: ring_simps) - -lemma less_add_iff1: - "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::pordered_ring))" -by (simp add: ring_simps) - -lemma less_add_iff2: - "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::pordered_ring))" -by (simp add: ring_simps) - -lemma le_add_iff1: - "(a*e + c \ b*e + d) = ((a-b)*e + c \ (d::'a::pordered_ring))" -by (simp add: ring_simps) +instance ordered_ring_strict \ ordered_ring by intro_classes -lemma le_add_iff2: - "(a*e + c \ b*e + d) = (c \ (b-a)*e + (d::'a::pordered_ring))" -by (simp add: ring_simps) - - -subsection {* Ordering Rules for Multiplication *} - -lemma mult_left_le_imp_le: - "[|c*a \ c*b; 0 < c|] ==> a \ (b::'a::ordered_semiring_strict)" -by (force simp add: mult_strict_left_mono linorder_not_less [symmetric]) - -lemma mult_right_le_imp_le: - "[|a*c \ b*c; 0 < c|] ==> a \ (b::'a::ordered_semiring_strict)" -by (force simp add: mult_strict_right_mono linorder_not_less [symmetric]) - -lemma mult_left_less_imp_less: - "[|c*a < c*b; 0 \ c|] ==> a < (b::'a::ordered_semiring)" -by (force simp add: mult_left_mono linorder_not_le [symmetric]) - -lemma mult_right_less_imp_less: - "[|a*c < b*c; 0 \ c|] ==> a < (b::'a::ordered_semiring)" -by (force simp add: mult_right_mono linorder_not_le [symmetric]) +context ordered_ring_strict +begin lemma mult_strict_left_mono_neg: - "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring_strict)" -apply (drule mult_strict_left_mono [of _ _ "-c"]) -apply (simp_all add: minus_mult_left [symmetric]) -done - -lemma mult_left_mono_neg: - "[|b \ a; c \ 0|] ==> c * a \ c * (b::'a::pordered_ring)" -apply (drule mult_left_mono [of _ _ "-c"]) -apply (simp_all add: minus_mult_left [symmetric]) -done + "b < a \ c < 0 \ c * a < c * b" + apply (drule mult_strict_left_mono [of _ _ "uminus c"]) + apply (simp_all add: minus_mult_left [symmetric]) + done lemma mult_strict_right_mono_neg: - "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring_strict)" -apply (drule mult_strict_right_mono [of _ _ "-c"]) -apply (simp_all add: minus_mult_right [symmetric]) -done - -lemma mult_right_mono_neg: - "[|b \ a; c \ 0|] ==> a * c \ (b::'a::pordered_ring) * c" -apply (drule mult_right_mono [of _ _ "-c"]) -apply (simp) -apply (simp_all add: minus_mult_right [symmetric]) -done - - -subsection{* Products of Signs *} - -lemma mult_pos_pos: "[| (0::'a::ordered_semiring_strict) < a; 0 < b |] ==> 0 < a*b" -by (drule mult_strict_left_mono [of 0 b], auto) - -lemma mult_nonneg_nonneg: "[| (0::'a::pordered_cancel_semiring) \ a; 0 \ b |] ==> 0 \ a*b" -by (drule mult_left_mono [of 0 b], auto) - -lemma mult_pos_neg: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> a*b < 0" -by (drule mult_strict_left_mono [of b 0], auto) - -lemma mult_nonneg_nonpos: "[| (0::'a::pordered_cancel_semiring) \ a; b \ 0 |] ==> a*b \ 0" -by (drule mult_left_mono [of b 0], auto) + "b < a \ c < 0 \ a * c < b * c" + apply (drule mult_strict_right_mono [of _ _ "uminus c"]) + apply (simp_all add: minus_mult_right [symmetric]) + done -lemma mult_pos_neg2: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> b*a < 0" -by (drule mult_strict_right_mono[of b 0], auto) - -lemma mult_nonneg_nonpos2: "[| (0::'a::pordered_cancel_semiring) \ a; b \ 0 |] ==> b*a \ 0" -by (drule mult_right_mono[of b 0], auto) - -lemma mult_neg_neg: "[| a < (0::'a::ordered_ring_strict); b < 0 |] ==> 0 < a*b" -by (drule mult_strict_right_mono_neg, auto) - -lemma mult_nonpos_nonpos: "[| a \ (0::'a::pordered_ring); b \ 0 |] ==> 0 \ a*b" -by (drule mult_right_mono_neg[of a 0 b ], auto) +lemma mult_neg_neg: + "a < 0 \ b < 0 \ 0 < a * b" + by (drule mult_strict_right_mono_neg, auto) -lemma zero_less_mult_pos: - "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)" -apply (cases "b\0") - apply (auto simp add: order_le_less linorder_not_less) -apply (drule_tac mult_pos_neg [of a b]) - apply (auto dest: order_less_not_sym) -done - -lemma zero_less_mult_pos2: - "[| 0 < b*a; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)" -apply (cases "b\0") - apply (auto simp add: order_le_less linorder_not_less) -apply (drule_tac mult_pos_neg2 [of a b]) - apply (auto dest: order_less_not_sym) -done +end lemma zero_less_mult_iff: - "((0::'a::ordered_ring_strict) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)" -apply (auto simp add: order_le_less linorder_not_less mult_pos_pos - mult_neg_neg) -apply (blast dest: zero_less_mult_pos) -apply (blast dest: zero_less_mult_pos2) -done - -lemma mult_eq_0_iff [simp]: - fixes a b :: "'a::ring_no_zero_divisors" - shows "(a * b = 0) = (a = 0 \ b = 0)" -by (cases "a = 0 \ b = 0", auto dest: no_zero_divisors) + fixes a :: "'a::ordered_ring_strict" + shows "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" + apply (auto simp add: le_less not_less mult_pos_pos mult_neg_neg) + apply (blast dest: zero_less_mult_pos) + apply (blast dest: zero_less_mult_pos2) + done instance ordered_ring_strict \ ring_no_zero_divisors apply intro_classes @@ -530,18 +592,57 @@ apply (force simp add: minus_mult_left[symmetric]) done -lemma split_mult_pos_le: "(0 \ a & 0 \ b) | (a \ 0 & b \ 0) \ 0 \ a * (b::_::pordered_ring)" -by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos) - -lemma split_mult_neg_le: "(0 \ a & b \ 0) | (a \ 0 & 0 \ b) \ a * b \ (0::_::pordered_cancel_semiring)" -by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) - lemma zero_le_square[simp]: "(0::'a::ordered_ring_strict) \ a*a" by (simp add: zero_le_mult_iff linorder_linear) lemma not_square_less_zero[simp]: "\ (a * a < (0::'a::ordered_ring_strict))" by (simp add: not_less) +text{*This list of rewrites simplifies ring terms by multiplying +everything out and bringing sums and products into a canonical form +(by ordered rewriting). As a result it decides ring equalities but +also helps with inequalities. *} +lemmas ring_simps = group_simps ring_distribs + + +class pordered_comm_ring = comm_ring + pordered_comm_semiring + +subclass (in pordered_comm_ring) pordered_ring by unfold_locales + +subclass (in pordered_comm_ring) pordered_cancel_comm_semiring by unfold_locales + +class ordered_semidom = comm_semiring_1_cancel + ordered_comm_semiring_strict + + (*previously ordered_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 + +end + +class ordered_idom = + comm_ring_1 + + ordered_comm_semiring_strict + + lordered_ab_group + + abs_if + sgn_if + (*previously ordered_ring*) + +instance ordered_idom \ ordered_ring_strict .. + +instance ordered_idom \ pordered_comm_ring .. + +class ordered_field = field + ordered_idom + +lemma linorder_neqE_ordered_idom: + fixes x y :: "'a :: ordered_idom" + assumes "x \ y" obtains "x < y" | "y < x" + using assms by (rule linorder_neqE) + +-- {* FIXME continue localization here *} + + text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semidom} theorems available to members of @{term ordered_idom} *} @@ -587,20 +688,6 @@ apply (blast intro: order_le_less_trans)+ done -lemma mult_mono: - "[|a \ b; c \ d; 0 \ b; 0 \ c|] - ==> a * c \ b * (d::'a::pordered_semiring)" -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::'a::pordered_semiring)" -apply (rule mult_mono) -apply (fast intro: order_trans)+ -done - lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semidom)" apply (insert mult_strict_mono [of 1 m 1 n]) apply (simp add: order_less_trans [OF zero_less_one]) @@ -805,43 +892,6 @@ mult_cancel_left1 mult_cancel_left2 -subsection {* Fields *} - -lemma right_inverse_eq: "b \ 0 ==> (a / b = 1) = (a = (b::'a::field))" -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::'a::field) = 1/a" -by (simp add: divide_inverse) - -lemma divide_self[simp]: "a \ 0 ==> a / (a::'a::field) = 1" - by (simp add: divide_inverse) - -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 add: divide_self) - -lemma divide_zero_left [simp]: "0/a = (0::'a::field)" -by (simp add: divide_inverse) - -lemma inverse_eq_divide: "inverse (a::'a::field) = 1/a" -by (simp add: divide_inverse) - -lemma add_divide_distrib: "(a+b)/(c::'a::field) = a/c + b/c" -by (simp add: divide_inverse ring_distribs) - (* what ordering?? this is a straight instance of mult_eq_0_iff text{*Compared with @{text mult_eq_0_iff}, this version removes the requirement of an ordering.*} @@ -1871,9 +1921,11 @@ lemma frac_le: "(0::'a::ordered_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 simp + apply (subst times_divide_eq_left) apply (rule mult_imp_le_div_pos, assumption) + thm mult_mono + thm mult_mono' apply (rule mult_mono) apply simp_all done