author | wenzelm |
Fri, 21 Mar 2025 22:26:18 +0100 | |
changeset 82317 | 231b6d8231c6 |
parent 80777 | 623d46973cbe |
permissions | -rw-r--r-- |
61546 | 1 |
(* Author: Steven Obua, TU Muenchen *) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
2 |
|
60500 | 3 |
section \<open>Various algebraic structures combined with a lattice\<close> |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
4 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
5 |
theory Lattice_Algebras |
65151 | 6 |
imports Complex_Main |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
7 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
8 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
9 |
class semilattice_inf_ab_group_add = ordered_ab_group_add + semilattice_inf |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
10 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
11 |
|
80777 | 12 |
lemma add_inf_distrib_left: "a + inf b c = inf (a + b) (a + c)" (is "?L=?R") |
13 |
proof (intro order.antisym) |
|
14 |
show "?R \<le> ?L" |
|
15 |
by (metis add_commute diff_le_eq inf_greatest inf_le1 inf_le2) |
|
16 |
qed simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
17 |
|
53240 | 18 |
lemma add_inf_distrib_right: "inf a b + c = inf (a + c) (b + c)" |
80777 | 19 |
using add_commute add_inf_distrib_left by presburger |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
20 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
21 |
end |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
22 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
23 |
class semilattice_sup_ab_group_add = ordered_ab_group_add + semilattice_sup |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
24 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
25 |
|
80777 | 26 |
lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a + c)" (is "?L = ?R") |
27 |
proof (rule order.antisym) |
|
28 |
show "?L \<le> ?R" |
|
29 |
by (metis add_commute le_diff_eq sup.bounded_iff sup_ge1 sup_ge2) |
|
30 |
qed simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
31 |
|
56228 | 32 |
lemma add_sup_distrib_right: "sup a b + c = sup (a + c) (b + c)" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
33 |
proof - |
56228 | 34 |
have "c + sup a b = sup (c+a) (c+b)" |
35 |
by (simp add: add_sup_distrib_left) |
|
36 |
then show ?thesis |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56228
diff
changeset
|
37 |
by (simp add: add.commute) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
38 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
39 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
40 |
end |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
41 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
42 |
class lattice_ab_group_add = ordered_ab_group_add + lattice |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
43 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
44 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
45 |
subclass semilattice_inf_ab_group_add .. |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
46 |
subclass semilattice_sup_ab_group_add .. |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
47 |
|
53240 | 48 |
lemmas add_sup_inf_distribs = |
49 |
add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
50 |
|
56228 | 51 |
lemma inf_eq_neg_sup: "inf a b = - sup (- a) (- b)" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
52 |
proof (rule inf_unique) |
53240 | 53 |
fix a b c :: 'a |
56228 | 54 |
show "- sup (- a) (- b) \<le> a" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
55 |
by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"]) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
56 |
(simp, simp add: add_sup_distrib_left) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
57 |
show "- sup (-a) (-b) \<le> b" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
58 |
by (rule add_le_imp_le_right [of _ "sup (uminus a) (uminus b)"]) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
59 |
(simp, simp add: add_sup_distrib_left) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
60 |
assume "a \<le> b" "a \<le> c" |
53240 | 61 |
then show "a \<le> - sup (-b) (-c)" |
62 |
by (subst neg_le_iff_le [symmetric]) (simp add: le_supI) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
63 |
qed |
53240 | 64 |
|
56228 | 65 |
lemma sup_eq_neg_inf: "sup a b = - inf (- a) (- b)" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
66 |
proof (rule sup_unique) |
53240 | 67 |
fix a b c :: 'a |
56228 | 68 |
show "a \<le> - inf (- a) (- b)" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
69 |
by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"]) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
70 |
(simp, simp add: add_inf_distrib_left) |
56228 | 71 |
show "b \<le> - inf (- a) (- b)" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
72 |
by (rule add_le_imp_le_right [of _ "inf (uminus a) (uminus b)"]) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
73 |
(simp, simp add: add_inf_distrib_left) |
65151 | 74 |
show "- inf (- a) (- b) \<le> c" if "a \<le> c" "b \<le> c" |
75 |
using that by (subst neg_le_iff_le [symmetric]) (simp add: le_infI) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
76 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
77 |
|
56228 | 78 |
lemma neg_inf_eq_sup: "- inf a b = sup (- a) (- b)" |
53240 | 79 |
by (simp add: inf_eq_neg_sup) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
80 |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
81 |
lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
82 |
using neg_inf_eq_sup [of b c, symmetric] by simp |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
83 |
|
56228 | 84 |
lemma neg_sup_eq_inf: "- sup a b = inf (- a) (- b)" |
53240 | 85 |
by (simp add: sup_eq_neg_inf) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
86 |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
87 |
lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
88 |
using neg_sup_eq_inf [of b c, symmetric] by simp |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
89 |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
90 |
lemma add_eq_inf_sup: "a + b = sup a b + inf a b" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
91 |
proof - |
56228 | 92 |
have "0 = - inf 0 (a - b) + inf (a - b) 0" |
53240 | 93 |
by (simp add: inf_commute) |
56228 | 94 |
then have "0 = sup 0 (b - a) + inf (a - b) 0" |
53240 | 95 |
by (simp add: inf_eq_neg_sup) |
56228 | 96 |
then have "0 = (- a + sup a b) + (inf a b + (- b))" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
97 |
by (simp only: add_sup_distrib_left add_inf_distrib_right) simp |
56228 | 98 |
then show ?thesis |
99 |
by (simp add: algebra_simps) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
100 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
101 |
|
53240 | 102 |
|
60500 | 103 |
subsection \<open>Positive Part, Negative Part, Absolute Value\<close> |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
104 |
|
53240 | 105 |
definition nprt :: "'a \<Rightarrow> 'a" |
106 |
where "nprt x = inf x 0" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
107 |
|
53240 | 108 |
definition pprt :: "'a \<Rightarrow> 'a" |
109 |
where "pprt x = sup x 0" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
110 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
111 |
lemma pprt_neg: "pprt (- x) = - nprt x" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
112 |
proof - |
56228 | 113 |
have "sup (- x) 0 = sup (- x) (- 0)" |
65151 | 114 |
by (simp only: minus_zero) |
56228 | 115 |
also have "\<dots> = - inf x 0" |
65151 | 116 |
by (simp only: neg_inf_eq_sup) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
117 |
finally have "sup (- x) 0 = - inf x 0" . |
56228 | 118 |
then show ?thesis |
65151 | 119 |
by (simp only: pprt_def nprt_def) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
120 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
121 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
122 |
lemma nprt_neg: "nprt (- x) = - pprt x" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
123 |
proof - |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
124 |
from pprt_neg have "pprt (- (- x)) = - nprt (- x)" . |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
125 |
then have "pprt x = - nprt (- x)" by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
126 |
then show ?thesis by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
127 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
128 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
129 |
lemma prts: "a = pprt a + nprt a" |
68406 | 130 |
by (simp add: pprt_def nprt_def flip: add_eq_inf_sup) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
131 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
132 |
lemma zero_le_pprt[simp]: "0 \<le> pprt a" |
53240 | 133 |
by (simp add: pprt_def) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
134 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
135 |
lemma nprt_le_zero[simp]: "nprt a \<le> 0" |
53240 | 136 |
by (simp add: nprt_def) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
137 |
|
60698 | 138 |
lemma le_eq_neg: "a \<le> - b \<longleftrightarrow> a + b \<le> 0" |
65151 | 139 |
(is "?lhs = ?rhs") |
53240 | 140 |
proof |
65151 | 141 |
assume ?lhs |
142 |
show ?rhs |
|
143 |
by (rule add_le_imp_le_right[of _ "uminus b" _]) (simp add: add.assoc \<open>?lhs\<close>) |
|
53240 | 144 |
next |
65151 | 145 |
assume ?rhs |
146 |
show ?lhs |
|
147 |
by (rule add_le_imp_le_right[of _ "b" _]) (simp add: \<open>?rhs\<close>) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
148 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
149 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
150 |
lemma pprt_0[simp]: "pprt 0 = 0" by (simp add: pprt_def) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
151 |
lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
152 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
153 |
lemma pprt_eq_id [simp, no_atp]: "0 \<le> x \<Longrightarrow> pprt x = x" |
46986 | 154 |
by (simp add: pprt_def sup_absorb1) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
155 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
156 |
lemma nprt_eq_id [simp, no_atp]: "x \<le> 0 \<Longrightarrow> nprt x = x" |
46986 | 157 |
by (simp add: nprt_def inf_absorb1) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
158 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
159 |
lemma pprt_eq_0 [simp, no_atp]: "x \<le> 0 \<Longrightarrow> pprt x = 0" |
46986 | 160 |
by (simp add: pprt_def sup_absorb2) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
161 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
162 |
lemma nprt_eq_0 [simp, no_atp]: "0 \<le> x \<Longrightarrow> nprt x = 0" |
46986 | 163 |
by (simp add: nprt_def inf_absorb2) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
164 |
|
60698 | 165 |
lemma sup_0_imp_0: |
166 |
assumes "sup a (- a) = 0" |
|
167 |
shows "a = 0" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
168 |
proof - |
65151 | 169 |
have pos: "0 \<le> a" if "sup a (- a) = 0" for a :: 'a |
60698 | 170 |
proof - |
171 |
from that have "sup a (- a) + a = a" |
|
56228 | 172 |
by simp |
173 |
then have "sup (a + a) 0 = a" |
|
174 |
by (simp add: add_sup_distrib_right) |
|
175 |
then have "sup (a + a) 0 \<le> a" |
|
176 |
by simp |
|
60698 | 177 |
then show ?thesis |
56228 | 178 |
by (blast intro: order_trans inf_sup_ord) |
60698 | 179 |
qed |
180 |
from assms have **: "sup (-a) (-(-a)) = 0" |
|
56228 | 181 |
by (simp add: sup_commute) |
65151 | 182 |
from pos[OF assms] pos[OF **] show "a = 0" |
56228 | 183 |
by simp |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
184 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
185 |
|
56228 | 186 |
lemma inf_0_imp_0: "inf a (- a) = 0 \<Longrightarrow> a = 0" |
80777 | 187 |
by (metis local.neg_0_equal_iff_equal neg_inf_eq_sup sup_0_imp_0) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
188 |
|
80777 | 189 |
lemma inf_0_eq_0 [simp]: "inf a (- a) = 0 \<longleftrightarrow> a = 0" |
190 |
by (metis inf_0_imp_0 inf.idem minus_zero) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
191 |
|
80777 | 192 |
lemma sup_0_eq_0 [simp]: "sup a (- a) = 0 \<longleftrightarrow> a = 0" |
193 |
by (metis minus_zero sup.idem sup_0_imp_0) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
194 |
|
60698 | 195 |
lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \<le> a + a \<longleftrightarrow> 0 \<le> a" |
196 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
197 |
proof |
60698 | 198 |
show ?rhs if ?lhs |
199 |
proof - |
|
200 |
from that have a: "inf (a + a) 0 = 0" |
|
201 |
by (simp add: inf_commute inf_absorb1) |
|
61546 | 202 |
have "inf a 0 + inf a 0 = inf (inf (a + a) 0) a" (is "?l = _") |
60698 | 203 |
by (simp add: add_sup_inf_distribs inf_aci) |
204 |
then have "?l = 0 + inf a 0" |
|
205 |
by (simp add: a, simp add: inf_commute) |
|
206 |
then have "inf a 0 = 0" |
|
207 |
by (simp only: add_right_cancel) |
|
208 |
then show ?thesis |
|
209 |
unfolding le_iff_inf by (simp add: inf_commute) |
|
210 |
qed |
|
211 |
show ?lhs if ?rhs |
|
212 |
by (simp add: add_mono[OF that that, simplified]) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
213 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
214 |
|
53240 | 215 |
lemma double_zero [simp]: "a + a = 0 \<longleftrightarrow> a = 0" |
73411 | 216 |
using add_nonneg_eq_0_iff order.eq_iff by auto |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
217 |
|
53240 | 218 |
lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \<longleftrightarrow> 0 < a" |
65151 | 219 |
by (meson le_less_trans less_add_same_cancel2 less_le_not_le |
220 |
zero_le_double_add_iff_zero_le_single_add) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
221 |
|
60698 | 222 |
lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \<le> 0 \<longleftrightarrow> a \<le> 0" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
223 |
proof - |
56228 | 224 |
have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)" |
60698 | 225 |
by (subst le_minus_iff) simp |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
226 |
moreover have "\<dots> \<longleftrightarrow> a \<le> 0" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
227 |
by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp |
56228 | 228 |
ultimately show ?thesis |
229 |
by blast |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
230 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
231 |
|
60698 | 232 |
lemma double_add_less_zero_iff_single_less_zero [simp]: "a + a < 0 \<longleftrightarrow> a < 0" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
233 |
proof - |
56228 | 234 |
have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)" |
235 |
by (subst less_minus_iff) simp |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
236 |
moreover have "\<dots> \<longleftrightarrow> a < 0" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
237 |
by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp |
56228 | 238 |
ultimately show ?thesis |
239 |
by blast |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
240 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
241 |
|
65151 | 242 |
declare neg_inf_eq_sup [simp] |
243 |
and neg_sup_eq_inf [simp] |
|
244 |
and diff_inf_eq_sup [simp] |
|
245 |
and diff_sup_eq_inf [simp] |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
246 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
247 |
lemma le_minus_self_iff: "a \<le> - a \<longleftrightarrow> a \<le> 0" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
248 |
proof - |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
249 |
from add_le_cancel_left [of "uminus a" "plus a a" zero] |
56228 | 250 |
have "a \<le> - a \<longleftrightarrow> a + a \<le> 0" |
68406 | 251 |
by (simp flip: add.assoc) |
56228 | 252 |
then show ?thesis |
253 |
by simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
254 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
255 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
256 |
lemma minus_le_self_iff: "- a \<le> a \<longleftrightarrow> 0 \<le> a" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
257 |
proof - |
56228 | 258 |
have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a" |
60698 | 259 |
using add_le_cancel_left [of "uminus a" zero "plus a a"] |
68406 | 260 |
by (simp flip: add.assoc) |
56228 | 261 |
then show ?thesis |
262 |
by simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
263 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
264 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
265 |
lemma zero_le_iff_zero_nprt: "0 \<le> a \<longleftrightarrow> nprt a = 0" |
53240 | 266 |
unfolding le_iff_inf by (simp add: nprt_def inf_commute) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
267 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
268 |
lemma le_zero_iff_zero_pprt: "a \<le> 0 \<longleftrightarrow> pprt a = 0" |
53240 | 269 |
unfolding le_iff_sup by (simp add: pprt_def sup_commute) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
270 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
271 |
lemma le_zero_iff_pprt_id: "0 \<le> a \<longleftrightarrow> pprt a = a" |
53240 | 272 |
unfolding le_iff_sup by (simp add: pprt_def sup_commute) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
273 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
274 |
lemma zero_le_iff_nprt_id: "a \<le> 0 \<longleftrightarrow> nprt a = a" |
53240 | 275 |
unfolding le_iff_inf by (simp add: nprt_def inf_commute) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
276 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
277 |
lemma pprt_mono [simp, no_atp]: "a \<le> b \<Longrightarrow> pprt a \<le> pprt b" |
53240 | 278 |
unfolding le_iff_sup by (simp add: pprt_def sup_aci sup_assoc [symmetric, of a]) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
279 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35040
diff
changeset
|
280 |
lemma nprt_mono [simp, no_atp]: "a \<le> b \<Longrightarrow> nprt a \<le> nprt b" |
53240 | 281 |
unfolding le_iff_inf by (simp add: nprt_def inf_aci inf_assoc [symmetric, of a]) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
282 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
283 |
end |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
284 |
|
56228 | 285 |
lemmas add_sup_inf_distribs = |
286 |
add_inf_distrib_right add_inf_distrib_left add_sup_distrib_right add_sup_distrib_left |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
287 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
288 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
289 |
class lattice_ab_group_add_abs = lattice_ab_group_add + abs + |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
290 |
assumes abs_lattice: "\<bar>a\<bar> = sup a (- a)" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
291 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
292 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
293 |
lemma abs_prts: "\<bar>a\<bar> = pprt a - nprt a" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
294 |
proof - |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
295 |
have "0 \<le> \<bar>a\<bar>" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
296 |
proof - |
56228 | 297 |
have a: "a \<le> \<bar>a\<bar>" and b: "- a \<le> \<bar>a\<bar>" |
298 |
by (auto simp add: abs_lattice) |
|
299 |
show ?thesis |
|
300 |
by (rule add_mono [OF a b, simplified]) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
301 |
qed |
56228 | 302 |
then have "0 \<le> sup a (- a)" |
303 |
unfolding abs_lattice . |
|
304 |
then have "sup (sup a (- a)) 0 = sup a (- a)" |
|
305 |
by (rule sup_absorb1) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
306 |
then show ?thesis |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
307 |
by (simp add: add_sup_inf_distribs ac_simps pprt_def nprt_def abs_lattice) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
308 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
309 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
310 |
subclass ordered_ab_group_add_abs |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
311 |
proof |
60698 | 312 |
have abs_ge_zero [simp]: "0 \<le> \<bar>a\<bar>" for a |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
313 |
proof - |
53240 | 314 |
have a: "a \<le> \<bar>a\<bar>" and b: "- a \<le> \<bar>a\<bar>" |
315 |
by (auto simp add: abs_lattice) |
|
316 |
show "0 \<le> \<bar>a\<bar>" |
|
317 |
by (rule add_mono [OF a b, simplified]) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
318 |
qed |
60698 | 319 |
have abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" for a b |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
320 |
by (simp add: abs_lattice le_supI) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
321 |
fix a b |
56228 | 322 |
show "0 \<le> \<bar>a\<bar>" |
323 |
by simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
324 |
show "a \<le> \<bar>a\<bar>" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
325 |
by (auto simp add: abs_lattice) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
326 |
show "\<bar>-a\<bar> = \<bar>a\<bar>" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
327 |
by (simp add: abs_lattice sup_commute) |
60698 | 328 |
show "- a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" if "a \<le> b" |
329 |
using that by (rule abs_leI) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
330 |
show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
331 |
proof - |
56228 | 332 |
have g: "\<bar>a\<bar> + \<bar>b\<bar> = sup (a + b) (sup (- a - b) (sup (- a + b) (a + (- b))))" |
60698 | 333 |
(is "_ = sup ?m ?n") |
57862 | 334 |
by (simp add: abs_lattice add_sup_inf_distribs ac_simps) |
56228 | 335 |
have a: "a + b \<le> sup ?m ?n" |
336 |
by simp |
|
337 |
have b: "- a - b \<le> ?n" |
|
338 |
by simp |
|
339 |
have c: "?n \<le> sup ?m ?n" |
|
340 |
by simp |
|
341 |
from b c have d: "- a - b \<le> sup ?m ?n" |
|
342 |
by (rule order_trans) |
|
343 |
have e: "- a - b = - (a + b)" |
|
344 |
by simp |
|
345 |
from a d e have "\<bar>a + b\<bar> \<le> sup ?m ?n" |
|
80777 | 346 |
by (metis abs_leI) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
347 |
with g[symmetric] show ?thesis by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
348 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
349 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
350 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
351 |
end |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
352 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
353 |
lemma sup_eq_if: |
60698 | 354 |
fixes a :: "'a::{lattice_ab_group_add,linorder}" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
355 |
shows "sup a (- a) = (if a < 0 then - a else a)" |
60698 | 356 |
using add_le_cancel_right [of a a "- a", symmetric, simplified] |
357 |
and add_le_cancel_right [of "-a" a a, symmetric, simplified] |
|
358 |
by (auto simp: sup_max max.absorb1 max.absorb2) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
359 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
360 |
lemma abs_if_lattice: |
60698 | 361 |
fixes a :: "'a::{lattice_ab_group_add_abs,linorder}" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
362 |
shows "\<bar>a\<bar> = (if a < 0 then - a else a)" |
53240 | 363 |
by auto |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
364 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
365 |
lemma estimate_by_abs: |
56228 | 366 |
fixes a b c :: "'a::lattice_ab_group_add_abs" |
60698 | 367 |
assumes "a + b \<le> c" |
368 |
shows "a \<le> c + \<bar>b\<bar>" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
369 |
proof - |
60698 | 370 |
from assms have "a \<le> c + (- b)" |
56228 | 371 |
by (simp add: algebra_simps) |
372 |
have "- b \<le> \<bar>b\<bar>" |
|
373 |
by (rule abs_ge_minus_self) |
|
374 |
then have "c + (- b) \<le> c + \<bar>b\<bar>" |
|
375 |
by (rule add_left_mono) |
|
60500 | 376 |
with \<open>a \<le> c + (- b)\<close> show ?thesis |
56228 | 377 |
by (rule order_trans) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
378 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
379 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
380 |
class lattice_ring = ordered_ring + lattice_ab_group_add_abs |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
381 |
begin |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
382 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
383 |
subclass semilattice_inf_ab_group_add .. |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
384 |
subclass semilattice_sup_ab_group_add .. |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
385 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
386 |
end |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
387 |
|
56228 | 388 |
lemma abs_le_mult: |
389 |
fixes a b :: "'a::lattice_ring" |
|
390 |
shows "\<bar>a * b\<bar> \<le> \<bar>a\<bar> * \<bar>b\<bar>" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
391 |
proof - |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
392 |
let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
393 |
let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" |
56228 | 394 |
have a: "\<bar>a\<bar> * \<bar>b\<bar> = ?x" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
395 |
by (simp only: abs_prts[of a] abs_prts[of b] algebra_simps) |
60698 | 396 |
have bh: "u = a \<Longrightarrow> v = b \<Longrightarrow> |
397 |
u * v = pprt a * pprt b + pprt a * nprt b + |
|
398 |
nprt a * pprt b + nprt a * nprt b" for u v :: 'a |
|
80777 | 399 |
by (metis add.commute combine_common_factor distrib_left prts) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
400 |
note b = this[OF refl[of a] refl[of b]] |
56228 | 401 |
have xy: "- ?x \<le> ?y" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
402 |
apply simp |
80777 | 403 |
by (meson add_increasing2 diff_le_eq neg_le_0_iff_le nprt_le_zero order.trans split_mult_pos_le zero_le_pprt) |
56228 | 404 |
have yx: "?y \<le> ?x" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53240
diff
changeset
|
405 |
apply simp |
80777 | 406 |
by (metis add_decreasing2 diff_0 diff_mono diff_zero mult_nonpos_nonneg mult_right_mono_neg mult_zero_left nprt_le_zero zero_le_pprt) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
407 |
show ?thesis |
80777 | 408 |
proof (rule abs_leI) |
409 |
show "a * b \<le> \<bar>a\<bar> * \<bar>b\<bar>" |
|
410 |
by (simp only: a b yx) |
|
411 |
show "- (a * b) \<le> \<bar>a\<bar> * \<bar>b\<bar>" |
|
412 |
by (metis a bh minus_le_iff xy) |
|
413 |
qed |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
414 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
415 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
416 |
instance lattice_ring \<subseteq> ordered_ring_abs |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
417 |
proof |
56228 | 418 |
fix a b :: "'a::lattice_ring" |
41528 | 419 |
assume a: "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)" |
56228 | 420 |
show "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
421 |
proof - |
56228 | 422 |
have s: "(0 \<le> a * b) \<or> (a * b \<le> 0)" |
80777 | 423 |
by (metis a split_mult_neg_le split_mult_pos_le) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
424 |
have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)" |
68406 | 425 |
by (simp flip: prts) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
426 |
show ?thesis |
56228 | 427 |
proof (cases "0 \<le> a * b") |
428 |
case True |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
429 |
then show ?thesis |
80777 | 430 |
using a split_mult_neg_le by fastforce |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
431 |
next |
56228 | 432 |
case False |
433 |
with s have "a * b \<le> 0" |
|
434 |
by simp |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
435 |
then show ?thesis |
80777 | 436 |
using a split_mult_pos_le by fastforce |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
437 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
438 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
439 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
440 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
441 |
lemma mult_le_prts: |
56228 | 442 |
fixes a b :: "'a::lattice_ring" |
443 |
assumes "a1 \<le> a" |
|
444 |
and "a \<le> a2" |
|
445 |
and "b1 \<le> b" |
|
446 |
and "b \<le> b2" |
|
447 |
shows "a * b \<le> |
|
53240 | 448 |
pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1" |
449 |
proof - |
|
450 |
have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" |
|
60698 | 451 |
by (subst prts[symmetric])+ simp |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
452 |
then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
453 |
by (simp add: algebra_simps) |
56228 | 454 |
moreover have "pprt a * pprt b \<le> pprt a2 * pprt b2" |
41528 | 455 |
by (simp_all add: assms mult_mono) |
56228 | 456 |
moreover have "pprt a * nprt b \<le> pprt a1 * nprt b2" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
457 |
proof - |
56228 | 458 |
have "pprt a * nprt b \<le> pprt a * nprt b2" |
41528 | 459 |
by (simp add: mult_left_mono assms) |
56228 | 460 |
moreover have "pprt a * nprt b2 \<le> pprt a1 * nprt b2" |
41528 | 461 |
by (simp add: mult_right_mono_neg assms) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
462 |
ultimately show ?thesis |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
463 |
by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
464 |
qed |
56228 | 465 |
moreover have "nprt a * pprt b \<le> nprt a2 * pprt b1" |
53240 | 466 |
proof - |
56228 | 467 |
have "nprt a * pprt b \<le> nprt a2 * pprt b" |
41528 | 468 |
by (simp add: mult_right_mono assms) |
56228 | 469 |
moreover have "nprt a2 * pprt b \<le> nprt a2 * pprt b1" |
41528 | 470 |
by (simp add: mult_left_mono_neg assms) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
471 |
ultimately show ?thesis |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
472 |
by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
473 |
qed |
56228 | 474 |
moreover have "nprt a * nprt b \<le> nprt a1 * nprt b1" |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
475 |
proof - |
56228 | 476 |
have "nprt a * nprt b \<le> nprt a * nprt b1" |
41528 | 477 |
by (simp add: mult_left_mono_neg assms) |
56228 | 478 |
moreover have "nprt a * nprt b1 \<le> nprt a1 * nprt b1" |
41528 | 479 |
by (simp add: mult_right_mono_neg assms) |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
480 |
ultimately show ?thesis |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
481 |
by simp |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
482 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
483 |
ultimately show ?thesis |
60698 | 484 |
by - (rule add_mono | simp)+ |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
485 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
486 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
487 |
lemma mult_ge_prts: |
56228 | 488 |
fixes a b :: "'a::lattice_ring" |
489 |
assumes "a1 \<le> a" |
|
490 |
and "a \<le> a2" |
|
491 |
and "b1 \<le> b" |
|
492 |
and "b \<le> b2" |
|
493 |
shows "a * b \<ge> |
|
53240 | 494 |
nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1" |
495 |
proof - |
|
56228 | 496 |
from assms have a1: "- a2 \<le> -a" |
53240 | 497 |
by auto |
56228 | 498 |
from assms have a2: "- a \<le> -a1" |
53240 | 499 |
by auto |
56228 | 500 |
from mult_le_prts[of "- a2" "- a" "- a1" "b1" b "b2", |
501 |
OF a1 a2 assms(3) assms(4), simplified nprt_neg pprt_neg] |
|
60698 | 502 |
have le: "- (a * b) \<le> |
503 |
- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + |
|
56228 | 504 |
- pprt a1 * pprt b1 + - pprt a2 * nprt b1" |
53240 | 505 |
by simp |
56228 | 506 |
then have "- (- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + |
507 |
- pprt a1 * pprt b1 + - pprt a2 * nprt b1) \<le> a * b" |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
508 |
by (simp only: minus_le_iff) |
56228 | 509 |
then show ?thesis |
510 |
by (simp add: algebra_simps) |
|
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
511 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
512 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
513 |
instance int :: lattice_ring |
53240 | 514 |
proof |
65151 | 515 |
show "\<bar>k\<bar> = sup k (- k)" for k :: int |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
516 |
by (auto simp add: sup_int_def) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
517 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
518 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
519 |
instance real :: lattice_ring |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
520 |
proof |
65151 | 521 |
show "\<bar>a\<bar> = sup a (- a)" for a :: real |
35040
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
522 |
by (auto simp add: sup_real_def) |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
523 |
qed |
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
524 |
|
e42e7f133d94
separate library theory for type classes combining lattices with various algebraic structures; c.f. cs. 7efe662e41b4
haftmann
parents:
diff
changeset
|
525 |
end |