1 (* Author: Steven Obua, TU Muenchen *)
3 section \<open>Various algebraic structures combined with a lattice\<close>
5 theory Lattice_Algebras
9 class semilattice_inf_ab_group_add = ordered_ab_group_add + semilattice_inf
12 lemma add_inf_distrib_left: "a + inf b c = inf (a + b) (a + c)"
14 apply (simp_all add: le_infI)
15 apply (rule add_le_imp_le_left [of "uminus a"])
16 apply (simp only: add.assoc [symmetric], simp add: diff_le_eq add.commute)
19 lemma add_inf_distrib_right: "inf a b + c = inf (a + c) (b + c)"
21 have "c + inf a b = inf (c + a) (c + b)"
22 by (simp add: add_inf_distrib_left)
24 by (simp add: add.commute)
29 class semilattice_sup_ab_group_add = ordered_ab_group_add + semilattice_sup
32 lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a + c)"
34 apply (rule add_le_imp_le_left [of "uminus a"])
35 apply (simp only: add.assoc [symmetric], simp)
36 apply (simp add: le_diff_eq add.commute)
38 apply (rule add_le_imp_le_left [of "a"], simp only: add.assoc[symmetric], simp)+
41 lemma add_sup_distrib_right: "sup a b + c = sup (a + c) (b + c)"
43 have "c + sup a b = sup (c+a) (c+b)"
44 by (simp add: add_sup_distrib_left)
46 by (simp add: add.commute)
51 class lattice_ab_group_add = ordered_ab_group_add + lattice
54 subclass semilattice_inf_ab_group_add ..
55 subclass semilattice_sup_ab_group_add ..
57 lemmas add_sup_inf_distribs =
58 add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left
60 lemma inf_eq_neg_sup: "inf a b = - sup (- a) (- b)"
61 proof (rule inf_unique)
63 show "- sup (- a) (- b) \<le> a"
64 by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"])
65 (simp, simp add: add_sup_distrib_left)
66 show "- sup (-a) (-b) \<le> b"
67 by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"])
68 (simp, simp add: add_sup_distrib_left)
69 assume "a \<le> b" "a \<le> c"
70 then show "a \<le> - sup (-b) (-c)"
71 by (subst neg_le_iff_le [symmetric]) (simp add: le_supI)
74 lemma sup_eq_neg_inf: "sup a b = - inf (- a) (- b)"
75 proof (rule sup_unique)
77 show "a \<le> - inf (- a) (- b)"
78 by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"])
79 (simp, simp add: add_inf_distrib_left)
80 show "b \<le> - inf (- a) (- b)"
81 by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"])
82 (simp, simp add: add_inf_distrib_left)
83 assume "a \<le> c" "b \<le> c"
84 then show "- inf (- a) (- b) \<le> c"
85 by (subst neg_le_iff_le [symmetric]) (simp add: le_infI)
88 lemma neg_inf_eq_sup: "- inf a b = sup (- a) (- b)"
89 by (simp add: inf_eq_neg_sup)
91 lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)"
92 using neg_inf_eq_sup [of b c, symmetric] by simp
94 lemma neg_sup_eq_inf: "- sup a b = inf (- a) (- b)"
95 by (simp add: sup_eq_neg_inf)
97 lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)"
98 using neg_sup_eq_inf [of b c, symmetric] by simp
100 lemma add_eq_inf_sup: "a + b = sup a b + inf a b"
102 have "0 = - inf 0 (a - b) + inf (a - b) 0"
103 by (simp add: inf_commute)
104 then have "0 = sup 0 (b - a) + inf (a - b) 0"
105 by (simp add: inf_eq_neg_sup)
106 then have "0 = (- a + sup a b) + (inf a b + (- b))"
107 by (simp only: add_sup_distrib_left add_inf_distrib_right) simp
109 by (simp add: algebra_simps)
113 subsection \<open>Positive Part, Negative Part, Absolute Value\<close>
115 definition nprt :: "'a \<Rightarrow> 'a"
116 where "nprt x = inf x 0"
118 definition pprt :: "'a \<Rightarrow> 'a"
119 where "pprt x = sup x 0"
121 lemma pprt_neg: "pprt (- x) = - nprt x"
123 have "sup (- x) 0 = sup (- x) (- 0)"
124 unfolding minus_zero ..
125 also have "\<dots> = - inf x 0"
126 unfolding neg_inf_eq_sup ..
127 finally have "sup (- x) 0 = - inf x 0" .
129 unfolding pprt_def nprt_def .
132 lemma nprt_neg: "nprt (- x) = - pprt x"
134 from pprt_neg have "pprt (- (- x)) = - nprt (- x)" .
135 then have "pprt x = - nprt (- x)" by simp
136 then show ?thesis by simp
139 lemma prts: "a = pprt a + nprt a"
140 by (simp add: pprt_def nprt_def add_eq_inf_sup[symmetric])
142 lemma zero_le_pprt[simp]: "0 \<le> pprt a"
143 by (simp add: pprt_def)
145 lemma nprt_le_zero[simp]: "nprt a \<le> 0"
146 by (simp add: nprt_def)
148 lemma le_eq_neg: "a \<le> - b \<longleftrightarrow> a + b \<le> 0"
154 apply (rule add_le_imp_le_right[of _ "uminus b" _])
155 apply (simp add: add.assoc)
161 apply (rule add_le_imp_le_right[of _ "b" _])
166 lemma pprt_0[simp]: "pprt 0 = 0" by (simp add: pprt_def)
167 lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def)
169 lemma pprt_eq_id [simp, no_atp]: "0 \<le> x \<Longrightarrow> pprt x = x"
170 by (simp add: pprt_def sup_absorb1)
172 lemma nprt_eq_id [simp, no_atp]: "x \<le> 0 \<Longrightarrow> nprt x = x"
173 by (simp add: nprt_def inf_absorb1)
175 lemma pprt_eq_0 [simp, no_atp]: "x \<le> 0 \<Longrightarrow> pprt x = 0"
176 by (simp add: pprt_def sup_absorb2)
178 lemma nprt_eq_0 [simp, no_atp]: "0 \<le> x \<Longrightarrow> nprt x = 0"
179 by (simp add: nprt_def inf_absorb2)
182 assumes "sup a (- a) = 0"
185 have p: "0 \<le> a" if "sup a (- a) = 0" for a :: 'a
187 from that have "sup a (- a) + a = a"
189 then have "sup (a + a) 0 = a"
190 by (simp add: add_sup_distrib_right)
191 then have "sup (a + a) 0 \<le> a"
194 by (blast intro: order_trans inf_sup_ord)
196 from assms have **: "sup (-a) (-(-a)) = 0"
197 by (simp add: sup_commute)
198 from p[OF assms] p[OF **] show "a = 0"
202 lemma inf_0_imp_0: "inf a (- a) = 0 \<Longrightarrow> a = 0"
203 apply (simp add: inf_eq_neg_sup)
204 apply (simp add: sup_commute)
205 apply (erule sup_0_imp_0)
208 lemma inf_0_eq_0 [simp, no_atp]: "inf a (- a) = 0 \<longleftrightarrow> a = 0"
210 apply (erule inf_0_imp_0)
214 lemma sup_0_eq_0 [simp, no_atp]: "sup a (- a) = 0 \<longleftrightarrow> a = 0"
216 apply (erule sup_0_imp_0)
220 lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
221 (is "?lhs \<longleftrightarrow> ?rhs")
225 from that have a: "inf (a + a) 0 = 0"
226 by (simp add: inf_commute inf_absorb1)
227 have "inf a 0 + inf a 0 = inf (inf (a + a) 0) a" (is "?l = _")
228 by (simp add: add_sup_inf_distribs inf_aci)
229 then have "?l = 0 + inf a 0"
230 by (simp add: a, simp add: inf_commute)
231 then have "inf a 0 = 0"
232 by (simp only: add_right_cancel)
234 unfolding le_iff_inf by (simp add: inf_commute)
237 by (simp add: add_mono[OF that that, simplified])
240 lemma double_zero [simp]: "a + a = 0 \<longleftrightarrow> a = 0"
241 (is "?lhs \<longleftrightarrow> ?rhs")
245 from that have "a + a + - a = - a"
247 then have "a + (a + - a) = - a"
248 by (simp only: add.assoc)
249 then have a: "- a = a"
253 apply (unfold neg_le_iff_le [symmetric, of a])
256 unfolding zero_le_double_add_iff_zero_le_single_add [symmetric, of a]
266 lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \<longleftrightarrow> 0 < a"
267 proof (cases "a = 0")
269 then show ?thesis by auto
280 unfolding double_zero [symmetric, of a]
285 lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
287 have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)"
288 by (subst le_minus_iff) simp
289 moreover have "\<dots> \<longleftrightarrow> a \<le> 0"
290 by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp
291 ultimately show ?thesis
295 lemma double_add_less_zero_iff_single_less_zero [simp]: "a + a < 0 \<longleftrightarrow> a < 0"
297 have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)"
298 by (subst less_minus_iff) simp
299 moreover have "\<dots> \<longleftrightarrow> a < 0"
300 by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp
301 ultimately show ?thesis
305 declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] diff_inf_eq_sup [simp] diff_sup_eq_inf [simp]
307 lemma le_minus_self_iff: "a \<le> - a \<longleftrightarrow> a \<le> 0"
309 from add_le_cancel_left [of "uminus a" "plus a a" zero]
310 have "a \<le> - a \<longleftrightarrow> a + a \<le> 0"
311 by (simp add: add.assoc[symmetric])
316 lemma minus_le_self_iff: "- a \<le> a \<longleftrightarrow> 0 \<le> a"
318 have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"
319 using add_le_cancel_left [of "uminus a" zero "plus a a"]
320 by (simp add: add.assoc[symmetric])
325 lemma zero_le_iff_zero_nprt: "0 \<le> a \<longleftrightarrow> nprt a = 0"
326 unfolding le_iff_inf by (simp add: nprt_def inf_commute)
328 lemma le_zero_iff_zero_pprt: "a \<le> 0 \<longleftrightarrow> pprt a = 0"
329 unfolding le_iff_sup by (simp add: pprt_def sup_commute)
331 lemma le_zero_iff_pprt_id: "0 \<le> a \<longleftrightarrow> pprt a = a"
332 unfolding le_iff_sup by (simp add: pprt_def sup_commute)
334 lemma zero_le_iff_nprt_id: "a \<le> 0 \<longleftrightarrow> nprt a = a"
335 unfolding le_iff_inf by (simp add: nprt_def inf_commute)
337 lemma pprt_mono [simp, no_atp]: "a \<le> b \<Longrightarrow> pprt a \<le> pprt b"
338 unfolding le_iff_sup by (simp add: pprt_def sup_aci sup_assoc [symmetric, of a])
340 lemma nprt_mono [simp, no_atp]: "a \<le> b \<Longrightarrow> nprt a \<le> nprt b"
341 unfolding le_iff_inf by (simp add: nprt_def inf_aci inf_assoc [symmetric, of a])
345 lemmas add_sup_inf_distribs =
346 add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left
349 class lattice_ab_group_add_abs = lattice_ab_group_add + abs +
350 assumes abs_lattice: "\<bar>a\<bar> = sup a (- a)"
353 lemma abs_prts: "\<bar>a\<bar> = pprt a - nprt a"
355 have "0 \<le> \<bar>a\<bar>"
357 have a: "a \<le> \<bar>a\<bar>" and b: "- a \<le> \<bar>a\<bar>"
358 by (auto simp add: abs_lattice)
360 by (rule add_mono [OF a b, simplified])
362 then have "0 \<le> sup a (- a)"
363 unfolding abs_lattice .
364 then have "sup (sup a (- a)) 0 = sup a (- a)"
365 by (rule sup_absorb1)
367 by (simp add: add_sup_inf_distribs ac_simps pprt_def nprt_def abs_lattice)
370 subclass ordered_ab_group_add_abs
372 have abs_ge_zero [simp]: "0 \<le> \<bar>a\<bar>" for a
374 have a: "a \<le> \<bar>a\<bar>" and b: "- a \<le> \<bar>a\<bar>"
375 by (auto simp add: abs_lattice)
376 show "0 \<le> \<bar>a\<bar>"
377 by (rule add_mono [OF a b, simplified])
379 have abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" for a b
380 by (simp add: abs_lattice le_supI)
382 show "0 \<le> \<bar>a\<bar>"
384 show "a \<le> \<bar>a\<bar>"
385 by (auto simp add: abs_lattice)
386 show "\<bar>-a\<bar> = \<bar>a\<bar>"
387 by (simp add: abs_lattice sup_commute)
388 show "- a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" if "a \<le> b"
389 using that by (rule abs_leI)
390 show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
392 have g: "\<bar>a\<bar> + \<bar>b\<bar> = sup (a + b) (sup (- a - b) (sup (- a + b) (a + (- b))))"
394 by (simp add: abs_lattice add_sup_inf_distribs ac_simps)
395 have a: "a + b \<le> sup ?m ?n"
397 have b: "- a - b \<le> ?n"
399 have c: "?n \<le> sup ?m ?n"
401 from b c have d: "- a - b \<le> sup ?m ?n"
402 by (rule order_trans)
403 have e: "- a - b = - (a + b)"
405 from a d e have "\<bar>a + b\<bar> \<le> sup ?m ?n"
407 apply (drule abs_leI)
408 apply (simp_all only: algebra_simps minus_add)
409 apply (metis add_uminus_conv_diff d sup_commute uminus_add_conv_diff)
411 with g[symmetric] show ?thesis by simp
418 fixes a :: "'a::{lattice_ab_group_add,linorder}"
419 shows "sup a (- a) = (if a < 0 then - a else a)"
420 using add_le_cancel_right [of a a "- a", symmetric, simplified]
421 and add_le_cancel_right [of "-a" a a, symmetric, simplified]
422 by (auto simp: sup_max max.absorb1 max.absorb2)
424 lemma abs_if_lattice:
425 fixes a :: "'a::{lattice_ab_group_add_abs,linorder}"
426 shows "\<bar>a\<bar> = (if a < 0 then - a else a)"
429 lemma estimate_by_abs:
430 fixes a b c :: "'a::lattice_ab_group_add_abs"
431 assumes "a + b \<le> c"
432 shows "a \<le> c + \<bar>b\<bar>"
434 from assms have "a \<le> c + (- b)"
435 by (simp add: algebra_simps)
436 have "- b \<le> \<bar>b\<bar>"
437 by (rule abs_ge_minus_self)
438 then have "c + (- b) \<le> c + \<bar>b\<bar>"
439 by (rule add_left_mono)
440 with \<open>a \<le> c + (- b)\<close> show ?thesis
441 by (rule order_trans)
444 class lattice_ring = ordered_ring + lattice_ab_group_add_abs
447 subclass semilattice_inf_ab_group_add ..
448 subclass semilattice_sup_ab_group_add ..
453 fixes a b :: "'a::lattice_ring"
454 shows "\<bar>a * b\<bar> \<le> \<bar>a\<bar> * \<bar>b\<bar>"
456 let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b"
457 let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
458 have a: "\<bar>a\<bar> * \<bar>b\<bar> = ?x"
459 by (simp only: abs_prts[of a] abs_prts[of b] algebra_simps)
460 have bh: "u = a \<Longrightarrow> v = b \<Longrightarrow>
461 u * v = pprt a * pprt b + pprt a * nprt b +
462 nprt a * pprt b + nprt a * nprt b" for u v :: 'a
463 apply (subst prts[of u], subst prts[of v])
464 apply (simp add: algebra_simps)
466 note b = this[OF refl[of a] refl[of b]]
467 have xy: "- ?x \<le> ?y"
469 apply (metis (full_types) add_increasing add_uminus_conv_diff
470 lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg
471 mult_nonpos_nonpos nprt_le_zero zero_le_pprt)
473 have yx: "?y \<le> ?x"
475 apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff
476 lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos
477 mult_nonpos_nonneg nprt_le_zero zero_le_pprt)
479 have i1: "a * b \<le> \<bar>a\<bar> * \<bar>b\<bar>"
480 by (simp only: a b yx)
481 have i2: "- (\<bar>a\<bar> * \<bar>b\<bar>) \<le> a * b"
482 by (simp only: a b xy)
486 apply (simp add: i2[simplified minus_le_iff])
490 instance lattice_ring \<subseteq> ordered_ring_abs
492 fix a b :: "'a::lattice_ring"
493 assume a: "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)"
494 show "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
496 have s: "(0 \<le> a * b) \<or> (a * b \<le> 0)"
498 apply (rule_tac split_mult_pos_le)
499 apply (rule_tac contrapos_np[of "a * b \<le> 0"])
501 apply (rule_tac split_mult_neg_le)
505 have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
506 by (simp add: prts[symmetric])
508 proof (cases "0 \<le> a * b")
511 apply (simp_all add: mulprts abs_prts)
513 apply (auto simp add:
515 iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt]
516 iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id])
517 apply(drule (1) mult_nonneg_nonpos[of a b], simp)
518 apply(drule (1) mult_nonneg_nonpos2[of b a], simp)
522 with s have "a * b \<le> 0"
525 apply (simp_all add: mulprts abs_prts)
527 apply (auto simp add: algebra_simps)
528 apply(drule (1) mult_nonneg_nonneg[of a b],simp)
529 apply(drule (1) mult_nonpos_nonpos[of a b],simp)
536 fixes a b :: "'a::lattice_ring"
542 pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1"
544 have "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
545 by (subst prts[symmetric])+ simp
546 then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
547 by (simp add: algebra_simps)
548 moreover have "pprt a * pprt b \<le> pprt a2 * pprt b2"
549 by (simp_all add: assms mult_mono)
550 moreover have "pprt a * nprt b \<le> pprt a1 * nprt b2"
552 have "pprt a * nprt b \<le> pprt a * nprt b2"
553 by (simp add: mult_left_mono assms)
554 moreover have "pprt a * nprt b2 \<le> pprt a1 * nprt b2"
555 by (simp add: mult_right_mono_neg assms)
556 ultimately show ?thesis
559 moreover have "nprt a * pprt b \<le> nprt a2 * pprt b1"
561 have "nprt a * pprt b \<le> nprt a2 * pprt b"
562 by (simp add: mult_right_mono assms)
563 moreover have "nprt a2 * pprt b \<le> nprt a2 * pprt b1"
564 by (simp add: mult_left_mono_neg assms)
565 ultimately show ?thesis
568 moreover have "nprt a * nprt b \<le> nprt a1 * nprt b1"
570 have "nprt a * nprt b \<le> nprt a * nprt b1"
571 by (simp add: mult_left_mono_neg assms)
572 moreover have "nprt a * nprt b1 \<le> nprt a1 * nprt b1"
573 by (simp add: mult_right_mono_neg assms)
574 ultimately show ?thesis
577 ultimately show ?thesis
578 by - (rule add_mono | simp)+
582 fixes a b :: "'a::lattice_ring"
588 nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1"
590 from assms have a1: "- a2 \<le> -a"
592 from assms have a2: "- a \<le> -a1"
594 from mult_le_prts[of "- a2" "- a" "- a1" "b1" b "b2",
595 OF a1 a2 assms(3) assms(4), simplified nprt_neg pprt_neg]
596 have le: "- (a * b) \<le>
597 - nprt a1 * pprt b2 + - nprt a2 * nprt b2 +
598 - pprt a1 * pprt b1 + - pprt a2 * nprt b1"
600 then have "- (- nprt a1 * pprt b2 + - nprt a2 * nprt b2 +
601 - pprt a1 * pprt b1 + - pprt a2 * nprt b1) \<le> a * b"
602 by (simp only: minus_le_iff)
604 by (simp add: algebra_simps)
607 instance int :: lattice_ring
610 show "\<bar>k\<bar> = sup k (- k)"
611 by (auto simp add: sup_int_def)
614 instance real :: lattice_ring
617 show "\<bar>a\<bar> = sup a (- a)"
618 by (auto simp add: sup_real_def)