# HG changeset patch # User wenzelm # Date 1395326329 -3600 # Node ID 0f6dc7512023496ee3a74dfc3be166ed3f14b9fd # Parent 00112abe9b256c9343ac82fa76cea7275cd78b05 tuned proofs; diff -r 00112abe9b25 -r 0f6dc7512023 src/HOL/Library/Lattice_Algebras.thy --- a/src/HOL/Library/Lattice_Algebras.thy Thu Mar 20 12:43:48 2014 +0000 +++ b/src/HOL/Library/Lattice_Algebras.thy Thu Mar 20 15:38:49 2014 +0100 @@ -18,9 +18,10 @@ 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)" + have "c + inf a b = inf (c + a) (c + b)" by (simp add: add_inf_distrib_left) - thus ?thesis by (simp add: add_commute) + then show ?thesis + by (simp add: add_commute) qed end @@ -37,10 +38,12 @@ apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+ done -lemma add_sup_distrib_right: "sup a b + c = sup (a+c) (b+c)" +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) + have "c + sup a b = sup (c+a) (c+b)" + by (simp add: add_sup_distrib_left) + then show ?thesis + by (simp add: add_commute) qed end @@ -54,10 +57,10 @@ 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)" +lemma inf_eq_neg_sup: "inf a b = - sup (- a) (- b)" proof (rule inf_unique) fix a b c :: 'a - show "- sup (-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) show "- sup (-a) (-b) \ b" @@ -68,26 +71,27 @@ by (subst neg_le_iff_le [symmetric]) (simp add: le_supI) qed -lemma sup_eq_neg_inf: "sup a b = - inf (-a) (-b)" +lemma sup_eq_neg_inf: "sup a b = - inf (- a) (- b)" proof (rule sup_unique) fix a b c :: 'a - show "a \ - inf (-a) (-b)" + 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) - show "b \ - inf (-a) (-b)" + 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) assume "a \ c" "b \ c" - then show "- inf (-a) (-b) \ c" by (subst neg_le_iff_le [symmetric]) (simp add: le_infI) + 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)" +lemma neg_inf_eq_sup: "- inf a b = sup (- a) (- b)" by (simp add: inf_eq_neg_sup) lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)" using neg_inf_eq_sup [of b c, symmetric] by simp -lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)" +lemma neg_sup_eq_inf: "- sup a b = inf (- a) (- b)" by (simp add: sup_eq_neg_inf) lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)" @@ -95,13 +99,14 @@ lemma add_eq_inf_sup: "a + b = sup a b + inf a b" proof - - have "0 = - inf 0 (a-b) + inf (a-b) 0" + 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" + then have "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))" + then have "0 = (- a + sup a b) + (inf a b + (- b))" by (simp only: add_sup_distrib_left add_inf_distrib_right) simp - then show ?thesis by (simp add: algebra_simps) + then show ?thesis + by (simp add: algebra_simps) qed @@ -115,10 +120,13 @@ 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 .. + 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 . + then show ?thesis + unfolding pprt_def nprt_def . qed lemma nprt_neg: "nprt (- x) = - pprt x" @@ -172,20 +180,26 @@ 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) + fix a :: 'a + assume hyp: "sup a (- a) = 0" + then have "sup a (- a) + a = a" + by simp + then have "sup (a + a) 0 = a" + by (simp add: add_sup_distrib_right) + then have "sup (a + a) 0 \ a" + by simp + then have "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 + then have 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" +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) @@ -206,24 +220,32 @@ 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=_") + assume "0 \ a + a" + then have 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) + then have "?l = 0 + inf a 0" + by (simp add: a, simp add: inf_commute) + then have "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]) + assume a: "0 \ a" + show "0 \ a + a" + by (simp add: add_mono[OF a a, simplified]) qed lemma double_zero [simp]: "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 + 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]) @@ -236,7 +258,8 @@ done next assume "a = 0" - then show "a + a = 0" by simp + then show "a + a = 0" + by simp qed lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \ 0 < a" @@ -261,19 +284,23 @@ 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) + have "a + a \ 0 \ 0 \ - (a + a)" + by (subst le_minus_iff, simp) moreover have "\ \ a \ 0" by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp - ultimately show ?thesis by blast + 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) + have "a + a < 0 \ 0 < - (a + a)" + by (subst less_minus_iff) simp moreover have "\ \ a < 0" by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp - ultimately show ?thesis by blast + ultimately show ?thesis + by blast qed declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] diff_inf_eq_sup [simp] diff_sup_eq_inf [simp] @@ -281,17 +308,19 @@ 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)" + have "a \ - a \ a + a \ 0" by (simp add: add_assoc[symmetric]) - thus ?thesis by simp + then show ?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)" + have "- a \ a \ 0 \ a + a" by (simp add: add_assoc[symmetric]) - thus ?thesis by simp + then show ?thesis + by simp qed lemma zero_le_iff_zero_nprt: "0 \ a \ nprt a = 0" @@ -314,7 +343,8 @@ end -lemmas add_sup_inf_distribs = add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left +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 + @@ -325,11 +355,15 @@ 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]) + 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 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 ac_simps pprt_def nprt_def abs_lattice) qed @@ -347,7 +381,8 @@ 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 "0 \ \a\" + by simp show "a \ \a\" by (auto simp add: abs_lattice) show "\-a\ = \a\" @@ -359,14 +394,20 @@ } 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") + have g: "\a\ + \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 ac_simps) - 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 - from a d e have "abs(a+b) <= sup ?m ?n" + 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 + from a d e have "\a + b\ \ sup ?m ?n" apply - apply (drule abs_leI) apply (simp_all only: algebra_simps ac_simps minus_add) @@ -379,7 +420,7 @@ end lemma sup_eq_if: - fixes a :: "'a\{lattice_ab_group_add, linorder}" + 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] @@ -388,18 +429,23 @@ qed lemma abs_if_lattice: - fixes a :: "'a\{lattice_ab_group_add_abs, linorder}" + fixes a :: "'a::{lattice_ab_group_add_abs, linorder}" shows "\a\ = (if a < 0 then - a else a)" by auto lemma estimate_by_abs: - "a + b <= (c::'a::lattice_ab_group_add_abs) \ a <= c + abs b" + fixes a b c :: "'a::lattice_ab_group_add_abs" + shows "a + b \ c \ a \ c + \b\" proof - - assume "a+b <= c" - then have "a <= c+(-b)" by (simp add: algebra_simps) - have "(-b) <= abs b" by (rule abs_ge_minus_self) - then have "c + (- b) \ c + \b\" by (rule add_left_mono) - with `a \ c + (- b)` show ?thesis by (rule order_trans) + assume "a + b \ c" + then have "a \ c + (- b)" + by (simp add: algebra_simps) + have "- b \ \b\" + by (rule abs_ge_minus_self) + then have "c + (- b) \ c + \b\" + by (rule add_left_mono) + with `a \ c + (- b)` show ?thesis + by (rule order_trans) qed class lattice_ring = ordered_ring + lattice_ab_group_add_abs @@ -410,15 +456,17 @@ end -lemma abs_le_mult: "abs (a * b) \ (abs a) * (abs (b::'a::lattice_ring))" +lemma abs_le_mult: + fixes a b :: "'a::lattice_ring" + shows "\a * b\ \ \a\ * \b\" 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" + have a: "\a\ * \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\ \ + 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]) @@ -426,16 +474,22 @@ done } note b = this[OF refl[of a] refl[of b]] - have xy: "- ?x <= ?y" + have xy: "- ?x \ ?y" apply simp - apply (metis (full_types) add_increasing add_uminus_conv_diff lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg mult_nonpos_nonpos nprt_le_zero zero_le_pprt) + apply (metis (full_types) add_increasing add_uminus_conv_diff + lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg + mult_nonpos_nonpos nprt_le_zero zero_le_pprt) done - have yx: "?y <= ?x" + have yx: "?y \ ?x" apply simp - apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos mult_nonpos_nonneg nprt_le_zero zero_le_pprt) + apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff + lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos + mult_nonpos_nonneg nprt_le_zero zero_le_pprt) 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) + have i1: "a * b \ \a\ * \b\" + by (simp only: a b yx) + have i2: "- (\a\ * \b\) \ a * b" + by (simp only: a b xy) show ?thesis apply (rule abs_leI) apply (simp add: i1) @@ -445,37 +499,38 @@ instance lattice_ring \ ordered_ring_abs proof - fix a b :: "'a\ lattice_ring" + fix a b :: "'a::lattice_ring" assume a: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0)" - show "abs (a*b) = abs a * abs b" + show "\a * b\ = \a\ * \b\" proof - - have s: "(0 <= a*b) | (a*b <= 0)" - apply (auto) + 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 contrapos_np[of "a * b \ 0"]) + apply simp apply (rule_tac split_mult_neg_le) - apply (insert a) - apply (blast) + using a + 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" + proof (cases "0 \ a * b") + case True then show ?thesis apply (simp_all add: mulprts abs_prts) - apply (insert a) + using a 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) + 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 + case False + with s have "a * b \ 0" + by simp then show ?thesis apply (simp_all add: mulprts abs_prts) apply (insert a) @@ -488,11 +543,12 @@ qed lemma mult_le_prts: - assumes "a1 <= (a::'a::lattice_ring)" - and "a <= a2" - and "b1 <= b" - and "b <= b2" - shows "a * b <= + fixes a b :: "'a::lattice_ring" + assumes "a1 \ a" + and "a \ a2" + and "b1 \ b" + and "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)" @@ -501,31 +557,31 @@ 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" + moreover have "pprt a * pprt b \ pprt a2 * pprt b2" by (simp_all add: assms mult_mono) - moreover have "pprt a * nprt b <= pprt a1 * nprt b2" + moreover have "pprt a * nprt b \ pprt a1 * nprt b2" proof - - have "pprt a * nprt b <= pprt a * nprt b2" + have "pprt a * nprt b \ pprt a * nprt b2" by (simp add: mult_left_mono assms) - moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" + moreover have "pprt a * nprt b2 \ pprt a1 * nprt b2" by (simp add: mult_right_mono_neg assms) ultimately show ?thesis by simp qed - moreover have "nprt a * pprt b <= nprt a2 * pprt b1" + moreover have "nprt a * pprt b \ nprt a2 * pprt b1" proof - - have "nprt a * pprt b <= nprt a2 * pprt b" + have "nprt a * pprt b \ nprt a2 * pprt b" by (simp add: mult_right_mono assms) - moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" + moreover have "nprt a2 * pprt b \ nprt a2 * pprt b1" by (simp add: mult_left_mono_neg assms) ultimately show ?thesis by simp qed - moreover have "nprt a * nprt b <= nprt a1 * nprt b1" + moreover have "nprt a * nprt b \ nprt a1 * nprt b1" proof - - have "nprt a * nprt b <= nprt a * nprt b1" + have "nprt a * nprt b \ nprt a * nprt b1" by (simp add: mult_left_mono_neg assms) - moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" + moreover have "nprt a * nprt b1 \ nprt a1 * nprt b1" by (simp add: mult_right_mono_neg assms) ultimately show ?thesis by simp @@ -537,36 +593,41 @@ qed lemma mult_ge_prts: - assumes "a1 <= (a::'a::lattice_ring)" - and "a <= a2" - and "b1 <= b" - and "b <= b2" - shows "a * b >= + fixes a b :: "'a::lattice_ring" + assumes "a1 \ a" + and "a \ a2" + and "b1 \ b" + and "b \ b2" + shows "a * b \ nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1" proof - - from assms have a1:"- a2 <= -a" + from assms have a1: "- a2 \ -a" by auto - from assms have a2: "-a <= -a1" + from assms have a2: "- a \ -a1" by auto - from mult_le_prts[of "-a2" "-a" "-a1" "b1" b "b2", OF a1 a2 assms(3) assms(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" + from mult_le_prts[of "- a2" "- a" "- a1" "b1" b "b2", + OF a1 a2 assms(3) assms(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" + 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 add: algebra_simps) + then show ?thesis + by (simp add: algebra_simps) qed instance int :: lattice_ring proof fix k :: int - show "abs k = sup k (- k)" + show "\k\ = sup k (- k)" by (auto simp add: sup_int_def) qed instance real :: lattice_ring proof fix a :: real - show "abs a = sup a (- a)" + show "\a\ = sup a (- a)" by (auto simp add: sup_real_def) qed