# HG changeset patch # User haftmann # Date 1265634512 -3600 # Node ID e682bb58707146c3ae956db651488d966eeb3328 # Parent 2ddc7edce107df0410488aae2e20ec4d1cab892d# Parent a1d93ce94235457a62cd2954c8070133e021a299 merged diff -r 2ddc7edce107 -r e682bb587071 NEWS --- a/NEWS Mon Feb 08 14:04:51 2010 +0100 +++ b/NEWS Mon Feb 08 14:08:32 2010 +0100 @@ -12,7 +12,7 @@ *** HOL *** -* more consistent naming of type classes involving orderings (and lattices): +* More consistent naming of type classes involving orderings (and lattices): lower_semilattice ~> semilattice_inf upper_semilattice ~> semilattice_sup @@ -33,12 +33,6 @@ pordered_ring_abs ~> ordered_ring_abs pordered_semiring ~> ordered_semiring - lordered_ab_group_add ~> lattice_ab_group_add - lordered_ab_group_add_abs ~> lattice_ab_group_add_abs - lordered_ab_group_add_meet ~> semilattice_inf_ab_group_add - lordered_ab_group_add_join ~> semilattice_sup_ab_group_add - lordered_ring ~> lattice_ring - ordered_ab_group_add ~> linordered_ab_group_add ordered_ab_semigroup_add ~> linordered_ab_semigroup_add ordered_cancel_ab_semigroup_add ~> linordered_cancel_ab_semigroup_add @@ -58,6 +52,15 @@ ordered_semiring_1_strict ~> linordered_semiring_1_strict ordered_semiring_strict ~> linordered_semiring_strict + The following slightly odd type classes have been moved to + a separate theory Library/Lattice_Algebras.thy: + + lordered_ab_group_add ~> lattice_ab_group_add + lordered_ab_group_add_abs ~> lattice_ab_group_add_abs + lordered_ab_group_add_meet ~> semilattice_inf_ab_group_add + lordered_ab_group_add_join ~> semilattice_sup_ab_group_add + lordered_ring ~> lattice_ring + INCOMPATIBILITY. * Index syntax for structures must be imported explicitly from library diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Feb 08 14:08:32 2010 +0100 @@ -2187,10 +2187,7 @@ {assume dc: "?c*?d < 0" from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0" - apply (simp add: mult_less_0_iff field_simps) - apply (rule add_neg_neg) - apply (simp_all add: mult_less_0_iff) - done + by (simp add: mult_less_0_iff field_simps) hence c:"?c \ 0" and d: "?d\ 0" by auto from add_frac_eq[OF c d, of "- ?t" "- ?s"] have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Finite_Set.thy Mon Feb 08 14:08:32 2010 +0100 @@ -3324,6 +3324,19 @@ end +context linordered_ab_group_add +begin + +lemma minus_Max_eq_Min [simp]: + "finite S \ S \ {} \ - (Max S) = Min (uminus ` S)" + by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min) + +lemma minus_Min_eq_Max [simp]: + "finite S \ S \ {} \ - (Min S) = Max (uminus ` S)" + by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max) + +end + subsection {* Expressing set operations via @{const fold} *} diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Int.thy --- a/src/HOL/Int.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Int.thy Mon Feb 08 14:08:32 2010 +0100 @@ -256,13 +256,6 @@ by (simp only: zsgn_def) qed -instance int :: lattice_ring -proof - fix k :: int - show "abs k = sup k (- k)" - by (auto simp add: sup_int_def zabs_def less_minus_self_iff [symmetric]) -qed - lemma zless_imp_add1_zle: "w < z \ w + (1\int) \ z" apply (cases w, cases z) apply (simp add: less le add One_int_def) diff -r 2ddc7edce107 -r e682bb587071 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/IsaMakefile Mon Feb 08 14:08:32 2010 +0100 @@ -384,6 +384,7 @@ Library/Permutations.thy Library/Bit.thy Library/FrechetDeriv.thy \ Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy \ Library/Inner_Product.thy Library/Kleene_Algebra.thy \ + Library/Lattice_Algebras.thy \ Library/Lattice_Syntax.thy Library/Library.thy \ Library/List_Prefix.thy Library/List_Set.thy Library/State_Monad.thy \ Library/Nat_Int_Bij.thy Library/Multiset.thy Library/Permutation.thy \ diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Library/Float.thy Mon Feb 08 14:08:32 2010 +0100 @@ -6,7 +6,7 @@ header {* Floating-Point Numbers *} theory Float -imports Complex_Main +imports Complex_Main Lattice_Algebras begin definition diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Library/Library.thy Mon Feb 08 14:08:32 2010 +0100 @@ -28,6 +28,7 @@ Fundamental_Theorem_Algebra Infinite_Set Inner_Product + Lattice_Algebras Lattice_Syntax ListVector Kleene_Algebra diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Matrix/ComputeFloat.thy --- a/src/HOL/Matrix/ComputeFloat.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Matrix/ComputeFloat.thy Mon Feb 08 14:08:32 2010 +0100 @@ -5,7 +5,7 @@ header {* Floating Point Representation of the Reals *} theory ComputeFloat -imports Complex_Main +imports Complex_Main Lattice_Algebras uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML") begin diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Matrix/LP.thy --- a/src/HOL/Matrix/LP.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Matrix/LP.thy Mon Feb 08 14:08:32 2010 +0100 @@ -3,7 +3,7 @@ *) theory LP -imports Main +imports Main Lattice_Algebras begin lemma linprog_dual_estimate: diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Matrix/Matrix.thy --- a/src/HOL/Matrix/Matrix.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Matrix/Matrix.thy Mon Feb 08 14:08:32 2010 +0100 @@ -3,7 +3,7 @@ *) theory Matrix -imports Main +imports Main Lattice_Algebras begin types 'a infmatrix = "nat \ nat \ 'a" diff -r 2ddc7edce107 -r e682bb587071 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/NSA/StarDef.thy Mon Feb 08 14:08:32 2010 +0100 @@ -849,12 +849,7 @@ simp add: abs_ge_self abs_leI abs_triangle_ineq)+ instance star :: (linordered_cancel_ab_semigroup_add) linordered_cancel_ab_semigroup_add .. -instance star :: (semilattice_inf_ab_group_add) semilattice_inf_ab_group_add .. -instance star :: (semilattice_inf_ab_group_add) semilattice_inf_ab_group_add .. -instance star :: (lattice_ab_group_add) lattice_ab_group_add .. -instance star :: (lattice_ab_group_add_abs) lattice_ab_group_add_abs -by (intro_classes, transfer, rule abs_lattice) subsection {* Ring and field classes *} @@ -934,7 +929,6 @@ instance star :: (ordered_ring) ordered_ring .. instance star :: (ordered_ring_abs) ordered_ring_abs by intro_classes (transfer, rule abs_eq_mult) -instance star :: (lattice_ring) lattice_ring .. instance star :: (abs_if) abs_if by (intro_classes, transfer, rule abs_if) diff -r 2ddc7edce107 -r e682bb587071 src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/OrderedGroup.thy Mon Feb 08 14:08:32 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 *} diff -r 2ddc7edce107 -r e682bb587071 src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/RealDef.thy Mon Feb 08 14:08:32 2010 +0100 @@ -426,8 +426,6 @@ by (simp only: real_sgn_def) qed -instance real :: lattice_ab_group_add .. - text{*The function @{term real_of_preal} requires many proofs, but it seems to be essential for proving completeness of the reals from that of the positive reals.*} @@ -1046,13 +1044,6 @@ lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \ abs(x + -l) + abs(y + -m)" by simp -instance real :: lattice_ring -proof - fix a::real - show "abs a = sup a (-a)" - by (auto simp add: real_abs_def sup_real_def) -qed - subsection {* Implementation of rational real numbers *} diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Ring_and_Field.thy Mon Feb 08 14:08:32 2010 +0100 @@ -2143,100 +2143,6 @@ assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" - -class lattice_ring = ordered_ring + lattice_ab_group_add_abs -begin - -subclass semilattice_inf_ab_group_add .. -subclass semilattice_sup_ab_group_add .. - -end - -lemma abs_le_mult: "abs (a * b) \ (abs a) * (abs (b::'a::lattice_ring))" -proof - - let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b" - let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" - have a: "(abs a) * (abs b) = ?x" - by (simp only: abs_prts[of a] abs_prts[of b] algebra_simps) - { - fix u v :: 'a - have bh: "\u = a; v = b\ \ - u * v = pprt a * pprt b + pprt a * nprt b + - nprt a * pprt b + nprt a * nprt b" - apply (subst prts[of u], subst prts[of v]) - apply (simp add: algebra_simps) - done - } - note b = this[OF refl[of a] refl[of b]] - note addm = add_mono[of "0::'a" _ "0::'a", simplified] - note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified] - have xy: "- ?x <= ?y" - apply (simp) - apply (rule_tac y="0::'a" in order_trans) - apply (rule addm2) - apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos) - apply (rule addm) - apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos) - done - have yx: "?y <= ?x" - apply (simp add:diff_def) - apply (rule_tac y=0 in order_trans) - apply (rule addm2, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+) - apply (rule addm, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+) - done - have i1: "a*b <= abs a * abs b" by (simp only: a b yx) - have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy) - show ?thesis - apply (rule abs_leI) - apply (simp add: i1) - apply (simp add: i2[simplified minus_le_iff]) - done -qed - -instance lattice_ring \ ordered_ring_abs -proof - fix a b :: "'a\ lattice_ring" - assume "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0)" - show "abs (a*b) = abs a * abs b" -proof - - have s: "(0 <= a*b) | (a*b <= 0)" - apply (auto) - apply (rule_tac split_mult_pos_le) - apply (rule_tac contrapos_np[of "a*b <= 0"]) - apply (simp) - apply (rule_tac split_mult_neg_le) - apply (insert prems) - apply (blast) - done - have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)" - by (simp add: prts[symmetric]) - show ?thesis - proof cases - assume "0 <= a * b" - then show ?thesis - apply (simp_all add: mulprts abs_prts) - apply (insert prems) - apply (auto simp add: - algebra_simps - iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt] - iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id]) - apply(drule (1) mult_nonneg_nonpos[of a b], simp) - apply(drule (1) mult_nonneg_nonpos2[of b a], simp) - done - next - assume "~(0 <= a*b)" - with s have "a*b <= 0" by simp - then show ?thesis - apply (simp_all add: mulprts abs_prts) - apply (insert prems) - apply (auto simp add: algebra_simps) - apply(drule (1) mult_nonneg_nonneg[of a b],simp) - apply(drule (1) mult_nonpos_nonpos[of a b],simp) - done - qed -qed -qed - context linordered_idom begin @@ -2308,76 +2214,6 @@ apply (simp add: order_less_imp_le) done - -subsection {* Bounds of products via negative and positive Part *} - -lemma mult_le_prts: - assumes - "a1 <= (a::'a::lattice_ring)" - "a <= a2" - "b1 <= b" - "b <= b2" - shows - "a * b <= pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1" -proof - - have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" - apply (subst prts[symmetric])+ - apply simp - done - then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" - by (simp add: algebra_simps) - moreover have "pprt a * pprt b <= pprt a2 * pprt b2" - by (simp_all add: prems mult_mono) - moreover have "pprt a * nprt b <= pprt a1 * nprt b2" - proof - - have "pprt a * nprt b <= pprt a * nprt b2" - by (simp add: mult_left_mono prems) - moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" - by (simp add: mult_right_mono_neg prems) - ultimately show ?thesis - by simp - qed - moreover have "nprt a * pprt b <= nprt a2 * pprt b1" - proof - - have "nprt a * pprt b <= nprt a2 * pprt b" - by (simp add: mult_right_mono prems) - moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" - by (simp add: mult_left_mono_neg prems) - ultimately show ?thesis - by simp - qed - moreover have "nprt a * nprt b <= nprt a1 * nprt b1" - proof - - have "nprt a * nprt b <= nprt a * nprt b1" - by (simp add: mult_left_mono_neg prems) - moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" - by (simp add: mult_right_mono_neg prems) - ultimately show ?thesis - by simp - qed - ultimately show ?thesis - by - (rule add_mono | simp)+ -qed - -lemma mult_ge_prts: - assumes - "a1 <= (a::'a::lattice_ring)" - "a <= a2" - "b1 <= b" - "b <= b2" - shows - "a * b >= nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1" -proof - - from prems have a1:"- a2 <= -a" by auto - from prems have a2: "-a <= -a1" by auto - from mult_le_prts[of "-a2" "-a" "-a1" "b1" b "b2", OF a1 a2 prems(3) prems(4), simplified nprt_neg pprt_neg] - have le: "- (a * b) <= - nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1" by simp - then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b" - by (simp only: minus_le_iff) - then show ?thesis by simp -qed - - code_modulename SML Ring_and_Field Arith diff -r 2ddc7edce107 -r e682bb587071 src/HOL/SupInf.thy --- a/src/HOL/SupInf.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/SupInf.thy Mon Feb 08 14:08:32 2010 +0100 @@ -6,38 +6,6 @@ imports RComplete begin -lemma minus_max_eq_min: - fixes x :: "'a::{lattice_ab_group_add, linorder}" - shows "- (max x y) = min (-x) (-y)" -by (metis le_imp_neg_le linorder_linear min_max.inf_absorb2 min_max.le_iff_inf min_max.le_iff_sup min_max.sup_absorb1) - -lemma minus_min_eq_max: - fixes x :: "'a::{lattice_ab_group_add, linorder}" - shows "- (min x y) = max (-x) (-y)" -by (metis minus_max_eq_min minus_minus) - -lemma minus_Max_eq_Min [simp]: - fixes S :: "'a::{lattice_ab_group_add, linorder} set" - shows "finite S \ S \ {} \ - (Max S) = Min (uminus ` S)" -proof (induct S rule: finite_ne_induct) - case (singleton x) - thus ?case by simp -next - case (insert x S) - thus ?case by (simp add: minus_max_eq_min) -qed - -lemma minus_Min_eq_Max [simp]: - fixes S :: "'a::{lattice_ab_group_add, linorder} set" - shows "finite S \ S \ {} \ - (Min S) = Max (uminus ` S)" -proof (induct S rule: finite_ne_induct) - case (singleton x) - thus ?case by simp -next - case (insert x S) - thus ?case by (simp add: minus_min_eq_max) -qed - instantiation real :: Sup begin definition diff -r 2ddc7edce107 -r e682bb587071 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Mon Feb 08 14:04:51 2010 +0100 +++ b/src/HOL/Transcendental.thy Mon Feb 08 14:08:32 2010 +0100 @@ -2904,10 +2904,12 @@ next case False hence "0 < \x\" and "- \x\ < \x\" by auto have "suminf (?c (-\x\)) - arctan (-\x\) = suminf (?c 0) - arctan 0" - by (rule suminf_eq_arctan_bounded[where x="0" and a="-\x\" and b="\x\", symmetric], auto simp add: `\x\ < r` `-\x\ < \x\`) + by (rule suminf_eq_arctan_bounded[where x="0" and a="-\x\" and b="\x\", symmetric]) + (simp_all only: `\x\ < r` `-\x\ < \x\` neg_less_iff_less) moreover have "suminf (?c x) - arctan x = suminf (?c (-\x\)) - arctan (-\x\)" - by (rule suminf_eq_arctan_bounded[where x="x" and a="-\x\" and b="\x\"], auto simp add: `\x\ < r` `-\x\ < \x\`) + by (rule suminf_eq_arctan_bounded[where x="x" and a="-\x\" and b="\x\"]) + (simp_all only: `\x\ < r` `-\x\ < \x\` neg_less_iff_less) ultimately show ?thesis using suminf_arctan_zero by auto qed