# HG changeset patch # User haftmann # Date 1265634411 -3600 # Node ID b8c8d01cc20d4bd4eb497854f488dbf17ec19900 # Parent 2c159d2cdae7787bc83fca9aba44340b747e504b separate library theory for type classes combining lattices with various algebraic structures; more simp rules diff -r 2c159d2cdae7 -r b8c8d01cc20d src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Mon Feb 08 14:06:48 2010 +0100 +++ b/src/HOL/OrderedGroup.thy Mon Feb 08 14:06:51 2010 +0100 @@ -710,7 +710,7 @@ subclass linordered_cancel_ab_semigroup_add .. -lemma neg_less_eq_nonneg: +lemma neg_less_eq_nonneg [simp]: "- a \ a \ 0 \ a" proof assume A: "- a \ a" show "0 \ a" @@ -728,8 +728,27 @@ show "0 \ a" using A . qed qed - -lemma less_eq_neg_nonpos: + +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" @@ -748,7 +767,7 @@ qed qed -lemma equal_neg_zero: +lemma equal_neg_zero [simp]: "a = - a \ a = 0" proof assume "a = 0" then show "a = - a" by simp @@ -765,9 +784,81 @@ qed qed -lemma neg_equal_zero: +lemma neg_equal_zero [simp]: "- a = a \ a = 0" - unfolding equal_neg_zero [symmetric] by auto + 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 @@ -941,375 +1032,6 @@ end - -subsection {* Lattice Ordered (Abelian) Groups *} - -class semilattice_inf_ab_group_add = ordered_ab_group_add + semilattice_inf -begin - -lemma add_inf_distrib_left: - "a + inf b c = inf (a + b) (a + c)" -apply (rule antisym) -apply (simp_all add: le_infI) -apply (rule add_le_imp_le_left [of "uminus a"]) -apply (simp only: add_assoc [symmetric], simp) -apply rule -apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp)+ -done - -lemma add_inf_distrib_right: - "inf a b + c = inf (a + c) (b + c)" -proof - - have "c + inf a b = inf (c+a) (c+b)" by (simp add: add_inf_distrib_left) - thus ?thesis by (simp add: add_commute) -qed - -end - -class semilattice_sup_ab_group_add = ordered_ab_group_add + semilattice_sup -begin - -lemma add_sup_distrib_left: - "a + sup b c = sup (a + b) (a + c)" -apply (rule antisym) -apply (rule add_le_imp_le_left [of "uminus a"]) -apply (simp only: add_assoc[symmetric], simp) -apply rule -apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+ -apply (rule le_supI) -apply (simp_all) -done - -lemma add_sup_distrib_right: - "sup a b + c = sup (a+c) (b+c)" -proof - - have "c + sup a b = sup (c+a) (c+b)" by (simp add: add_sup_distrib_left) - thus ?thesis by (simp add: add_commute) -qed - -end - -class lattice_ab_group_add = ordered_ab_group_add + lattice -begin - -subclass semilattice_inf_ab_group_add .. -subclass semilattice_sup_ab_group_add .. - -lemmas add_sup_inf_distribs = add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left - -lemma inf_eq_neg_sup: "inf a b = - sup (-a) (-b)" -proof (rule inf_unique) - fix a b :: 'a - show "- sup (-a) (-b) \ a" - by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"]) - (simp, simp add: add_sup_distrib_left) -next - fix a b :: 'a - show "- sup (-a) (-b) \ b" - by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"]) - (simp, simp add: add_sup_distrib_left) -next - fix a b c :: 'a - assume "a \ b" "a \ c" - then show "a \ - sup (-b) (-c)" by (subst neg_le_iff_le [symmetric]) - (simp add: le_supI) -qed - -lemma sup_eq_neg_inf: "sup a b = - inf (-a) (-b)" -proof (rule sup_unique) - fix a b :: 'a - show "a \ - inf (-a) (-b)" - by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"]) - (simp, simp add: add_inf_distrib_left) -next - fix a b :: 'a - show "b \ - inf (-a) (-b)" - by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"]) - (simp, simp add: add_inf_distrib_left) -next - fix a b c :: 'a - assume "a \ c" "b \ c" - then show "- inf (-a) (-b) \ c" by (subst neg_le_iff_le [symmetric]) - (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) - hence "0 = sup 0 (b-a) + inf (a-b) 0" by (simp add: inf_eq_neg_sup) - hence "0 = (-a + sup a b) + (inf a b + (-b))" - by (simp add: add_sup_distrib_left add_inf_distrib_right) - (simp add: algebra_simps) - thus ?thesis by (simp add: algebra_simps) -qed - -subsection {* Positive Part, Negative Part, Absolute Value *} - -definition - nprt :: "'a \ 'a" where - "nprt x = inf x 0" - -definition - 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]) - -lemma zero_le_pprt[simp]: "0 \ pprt a" -by (simp add: pprt_def) - -lemma nprt_le_zero[simp]: "nprt a \ 0" -by (simp add: nprt_def) - -lemma le_eq_neg: "a \ - b \ a + b \ 0" (is "?l = ?r") -proof - - have a: "?l \ ?r" - apply (auto) - apply (rule add_le_imp_le_right[of _ "uminus b" _]) - apply (simp add: add_assoc) - done - have b: "?r \ ?l" - apply (auto) - apply (rule add_le_imp_le_right[of _ "b" _]) - apply (simp) - done - from a b show ?thesis by blast -qed - -lemma pprt_0[simp]: "pprt 0 = 0" by (simp add: pprt_def) -lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def) - -lemma pprt_eq_id [simp, noatp]: "0 \ x \ pprt x = x" - by (simp add: pprt_def sup_aci sup_absorb1) - -lemma nprt_eq_id [simp, noatp]: "x \ 0 \ nprt x = x" - by (simp add: nprt_def inf_aci inf_absorb1) - -lemma pprt_eq_0 [simp, noatp]: "x \ 0 \ pprt x = 0" - by (simp add: pprt_def sup_aci sup_absorb2) - -lemma nprt_eq_0 [simp, noatp]: "0 \ x \ nprt x = 0" - by (simp add: nprt_def inf_aci inf_absorb2) - -lemma sup_0_imp_0: "sup a (- a) = 0 \ a = 0" -proof - - { - fix a::'a - assume hyp: "sup a (-a) = 0" - hence "sup a (-a) + a = a" by (simp) - hence "sup (a+a) 0 = a" by (simp add: add_sup_distrib_right) - hence "sup (a+a) 0 <= a" by (simp) - hence "0 <= a" by (blast intro: order_trans inf_sup_ord) - } - note p = this - assume hyp:"sup a (-a) = 0" - hence hyp2:"sup (-a) (-(-a)) = 0" by (simp add: sup_commute) - from p[OF hyp] p[OF hyp2] show "a = 0" by simp -qed - -lemma inf_0_imp_0: "inf a (-a) = 0 \ a = 0" -apply (simp add: inf_eq_neg_sup) -apply (simp add: sup_commute) -apply (erule sup_0_imp_0) -done - -lemma inf_0_eq_0 [simp, noatp]: "inf a (- a) = 0 \ a = 0" -by (rule, erule inf_0_imp_0) simp - -lemma sup_0_eq_0 [simp, noatp]: "sup a (- a) = 0 \ a = 0" -by (rule, erule sup_0_imp_0) simp - -lemma zero_le_double_add_iff_zero_le_single_add [simp]: - "0 \ a + a \ 0 \ a" -proof - assume "0 <= a + a" - hence a:"inf (a+a) 0 = 0" by (simp add: inf_commute inf_absorb1) - have "(inf a 0)+(inf a 0) = inf (inf (a+a) 0) a" (is "?l=_") - by (simp add: add_sup_inf_distribs inf_aci) - hence "?l = 0 + inf a 0" by (simp add: a, simp add: inf_commute) - hence "inf a 0 = 0" by (simp only: add_right_cancel) - then show "0 <= a" unfolding le_iff_inf by (simp add: inf_commute) -next - assume a: "0 <= a" - show "0 <= a + a" by (simp add: add_mono[OF a a, simplified]) -qed - -lemma double_zero: "a + a = 0 \ a = 0" -proof - assume assm: "a + a = 0" - then have "a + a + - a = - a" by simp - then have "a + (a + - a) = - a" by (simp only: add_assoc) - then have a: "- a = a" by simp - show "a = 0" apply (rule antisym) - apply (unfold neg_le_iff_le [symmetric, of a]) - unfolding a apply simp - unfolding zero_le_double_add_iff_zero_le_single_add [symmetric, of a] - unfolding assm unfolding le_less apply simp_all done -next - assume "a = 0" then show "a + a = 0" by simp -qed - -lemma zero_less_double_add_iff_zero_less_single_add: - "0 < a + a \ 0 < a" -proof (cases "a = 0") - case True then show ?thesis by auto -next - case False then show ?thesis (*FIXME tune proof*) - unfolding less_le apply simp apply rule - apply clarify - apply rule - apply assumption - apply (rule notI) - unfolding double_zero [symmetric, of a] apply simp - done -qed - -lemma double_add_le_zero_iff_single_add_le_zero [simp]: - "a + a \ 0 \ a \ 0" -proof - - have "a + a \ 0 \ 0 \ - (a + a)" by (subst le_minus_iff, simp) - moreover have "\ \ a \ 0" by (simp add: zero_le_double_add_iff_zero_le_single_add) - ultimately show ?thesis by blast -qed - -lemma double_add_less_zero_iff_single_less_zero [simp]: - "a + a < 0 \ a < 0" -proof - - have "a + a < 0 \ 0 < - (a + a)" by (subst less_minus_iff, simp) - moreover have "\ \ a < 0" by (simp add: zero_less_double_add_iff_zero_less_single_add) - 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" -unfolding le_iff_inf by (simp add: nprt_def inf_commute) - -lemma le_zero_iff_zero_pprt: "a \ 0 \ pprt a = 0" -unfolding le_iff_sup by (simp add: pprt_def sup_commute) - -lemma le_zero_iff_pprt_id: "0 \ a \ pprt a = a" -unfolding le_iff_sup by (simp add: pprt_def sup_commute) - -lemma zero_le_iff_nprt_id: "a \ 0 \ nprt a = a" -unfolding le_iff_inf by (simp add: nprt_def inf_commute) - -lemma pprt_mono [simp, noatp]: "a \ b \ pprt a \ pprt b" -unfolding le_iff_sup by (simp add: pprt_def sup_aci sup_assoc [symmetric, of a]) - -lemma nprt_mono [simp, noatp]: "a \ b \ nprt a \ nprt b" -unfolding le_iff_inf by (simp add: 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 lattice_ab_group_add_abs = lattice_ab_group_add + 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 ordered_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) - fix a b - show "0 \ \a\" by simp - show "a \ \a\" - by (auto simp add: abs_lattice) - show "\-a\ = \a\" - by (simp add: abs_lattice sup_commute) - show "a \ b \ - a \ b \ \a\ \ b" by (fact abs_leI) - 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 - -end - -lemma sup_eq_if: - fixes a :: "'a\{lattice_ab_group_add, linorder}" - shows "sup a (- a) = (if a < 0 then - a else a)" -proof - - note add_le_cancel_right [of a a "- a", symmetric, simplified] - moreover note add_le_cancel_right [of "-a" a a, symmetric, simplified] - then show ?thesis by (auto simp: sup_max min_max.sup_absorb1 min_max.sup_absorb2) -qed - -lemma abs_if_lattice: - fixes a :: "'a\{lattice_ab_group_add_abs, linorder}" - shows "\a\ = (if a < 0 then - a else a)" -by auto - - text {* Needed for abelian cancellation simprocs: *} lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)" @@ -1346,14 +1068,6 @@ apply (simp_all add: prems) done -lemma estimate_by_abs: - "a + b <= (c::'a::lattice_ab_group_add_abs) \ a <= c + abs b" -proof - - assume "a+b <= c" - hence 2: "a <= c+(-b)" by (simp add: algebra_simps) - have 3: "(-b) <= abs b" by (rule abs_ge_minus_self) - show ?thesis by (rule le_add_right_mono[OF 2 3]) -qed subsection {* Tools setup *}