author | paulson |
Tue, 12 Jan 2010 16:55:59 +0000 | |
changeset 34883 | 77f0d11dec76 |
parent 34209 | c7f621786035 |
child 34973 | ae634fad947e |
child 36092 | 8f1e60d9f7cc |
permissions | -rw-r--r-- |
21249 | 1 |
(* Title: HOL/Lattices.thy |
2 |
Author: Tobias Nipkow |
|
3 |
*) |
|
4 |
||
22454 | 5 |
header {* Abstract lattices *} |
21249 | 6 |
|
7 |
theory Lattices |
|
30302 | 8 |
imports Orderings |
21249 | 9 |
begin |
10 |
||
28562 | 11 |
subsection {* Lattices *} |
21249 | 12 |
|
25206 | 13 |
notation |
25382 | 14 |
less_eq (infix "\<sqsubseteq>" 50) and |
32568 | 15 |
less (infix "\<sqsubset>" 50) and |
16 |
top ("\<top>") and |
|
17 |
bot ("\<bottom>") |
|
25206 | 18 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
19 |
class lower_semilattice = order + |
21249 | 20 |
fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70) |
22737 | 21 |
assumes inf_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x" |
22 |
and inf_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y" |
|
21733 | 23 |
and inf_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z" |
21249 | 24 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
25 |
class upper_semilattice = order + |
21249 | 26 |
fixes sup :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) |
22737 | 27 |
assumes sup_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y" |
28 |
and sup_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y" |
|
21733 | 29 |
and sup_least: "y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<squnion> z \<sqsubseteq> x" |
26014 | 30 |
begin |
31 |
||
32 |
text {* Dual lattice *} |
|
33 |
||
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
34 |
lemma dual_semilattice: |
26014 | 35 |
"lower_semilattice (op \<ge>) (op >) sup" |
27682 | 36 |
by (rule lower_semilattice.intro, rule dual_order) |
37 |
(unfold_locales, simp_all add: sup_least) |
|
26014 | 38 |
|
39 |
end |
|
21249 | 40 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
41 |
class lattice = lower_semilattice + upper_semilattice |
21249 | 42 |
|
25382 | 43 |
|
28562 | 44 |
subsubsection {* Intro and elim rules*} |
21733 | 45 |
|
46 |
context lower_semilattice |
|
47 |
begin |
|
21249 | 48 |
|
32064 | 49 |
lemma le_infI1: |
50 |
"a \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" |
|
51 |
by (rule order_trans) auto |
|
21249 | 52 |
|
32064 | 53 |
lemma le_infI2: |
54 |
"b \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" |
|
55 |
by (rule order_trans) auto |
|
21733 | 56 |
|
32064 | 57 |
lemma le_infI: "x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<sqinter> b" |
58 |
by (blast intro: inf_greatest) |
|
21249 | 59 |
|
32064 | 60 |
lemma le_infE: "x \<sqsubseteq> a \<sqinter> b \<Longrightarrow> (x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> P) \<Longrightarrow> P" |
61 |
by (blast intro: order_trans le_infI1 le_infI2) |
|
21249 | 62 |
|
21734 | 63 |
lemma le_inf_iff [simp]: |
32064 | 64 |
"x \<sqsubseteq> y \<sqinter> z \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<sqsubseteq> z" |
65 |
by (blast intro: le_infI elim: le_infE) |
|
21733 | 66 |
|
32064 | 67 |
lemma le_iff_inf: |
68 |
"x \<sqsubseteq> y \<longleftrightarrow> x \<sqinter> y = x" |
|
69 |
by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1]) |
|
21249 | 70 |
|
25206 | 71 |
lemma mono_inf: |
72 |
fixes f :: "'a \<Rightarrow> 'b\<Colon>lower_semilattice" |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
73 |
shows "mono f \<Longrightarrow> f (A \<sqinter> B) \<sqsubseteq> f A \<sqinter> f B" |
25206 | 74 |
by (auto simp add: mono_def intro: Lattices.inf_greatest) |
21733 | 75 |
|
25206 | 76 |
end |
21733 | 77 |
|
78 |
context upper_semilattice |
|
79 |
begin |
|
21249 | 80 |
|
32064 | 81 |
lemma le_supI1: |
82 |
"x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" |
|
25062 | 83 |
by (rule order_trans) auto |
21249 | 84 |
|
32064 | 85 |
lemma le_supI2: |
86 |
"x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" |
|
25062 | 87 |
by (rule order_trans) auto |
21733 | 88 |
|
32064 | 89 |
lemma le_supI: |
90 |
"a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> a \<squnion> b \<sqsubseteq> x" |
|
26014 | 91 |
by (blast intro: sup_least) |
21249 | 92 |
|
32064 | 93 |
lemma le_supE: |
94 |
"a \<squnion> b \<sqsubseteq> x \<Longrightarrow> (a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> P) \<Longrightarrow> P" |
|
95 |
by (blast intro: le_supI1 le_supI2 order_trans) |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
96 |
|
32064 | 97 |
lemma le_sup_iff [simp]: |
98 |
"x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z" |
|
99 |
by (blast intro: le_supI elim: le_supE) |
|
21733 | 100 |
|
32064 | 101 |
lemma le_iff_sup: |
102 |
"x \<sqsubseteq> y \<longleftrightarrow> x \<squnion> y = y" |
|
103 |
by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1]) |
|
21734 | 104 |
|
25206 | 105 |
lemma mono_sup: |
106 |
fixes f :: "'a \<Rightarrow> 'b\<Colon>upper_semilattice" |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
107 |
shows "mono f \<Longrightarrow> f A \<squnion> f B \<sqsubseteq> f (A \<squnion> B)" |
25206 | 108 |
by (auto simp add: mono_def intro: Lattices.sup_least) |
21733 | 109 |
|
25206 | 110 |
end |
23878 | 111 |
|
21733 | 112 |
|
32064 | 113 |
subsubsection {* Equational laws *} |
21249 | 114 |
|
21733 | 115 |
context lower_semilattice |
116 |
begin |
|
117 |
||
118 |
lemma inf_commute: "(x \<sqinter> y) = (y \<sqinter> x)" |
|
32064 | 119 |
by (rule antisym) auto |
21733 | 120 |
|
121 |
lemma inf_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)" |
|
32064 | 122 |
by (rule antisym) (auto intro: le_infI1 le_infI2) |
21733 | 123 |
|
124 |
lemma inf_idem[simp]: "x \<sqinter> x = x" |
|
32064 | 125 |
by (rule antisym) auto |
21733 | 126 |
|
127 |
lemma inf_left_idem[simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y" |
|
32064 | 128 |
by (rule antisym) (auto intro: le_infI2) |
21733 | 129 |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
130 |
lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x" |
32064 | 131 |
by (rule antisym) auto |
21733 | 132 |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
133 |
lemma inf_absorb2: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y" |
32064 | 134 |
by (rule antisym) auto |
21733 | 135 |
|
136 |
lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)" |
|
32064 | 137 |
by (rule mk_left_commute [of inf]) (fact inf_assoc inf_commute)+ |
138 |
||
139 |
lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem |
|
21733 | 140 |
|
141 |
end |
|
142 |
||
143 |
context upper_semilattice |
|
144 |
begin |
|
21249 | 145 |
|
21733 | 146 |
lemma sup_commute: "(x \<squnion> y) = (y \<squnion> x)" |
32064 | 147 |
by (rule antisym) auto |
21733 | 148 |
|
149 |
lemma sup_assoc: "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)" |
|
32064 | 150 |
by (rule antisym) (auto intro: le_supI1 le_supI2) |
21733 | 151 |
|
152 |
lemma sup_idem[simp]: "x \<squnion> x = x" |
|
32064 | 153 |
by (rule antisym) auto |
21733 | 154 |
|
155 |
lemma sup_left_idem[simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y" |
|
32064 | 156 |
by (rule antisym) (auto intro: le_supI2) |
21733 | 157 |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
158 |
lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x" |
32064 | 159 |
by (rule antisym) auto |
21733 | 160 |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
161 |
lemma sup_absorb2: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y" |
32064 | 162 |
by (rule antisym) auto |
21249 | 163 |
|
21733 | 164 |
lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)" |
32064 | 165 |
by (rule mk_left_commute [of sup]) (fact sup_assoc sup_commute)+ |
21733 | 166 |
|
32064 | 167 |
lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem |
21733 | 168 |
|
169 |
end |
|
21249 | 170 |
|
21733 | 171 |
context lattice |
172 |
begin |
|
173 |
||
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
174 |
lemma dual_lattice: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
175 |
"lattice (op \<ge>) (op >) sup inf" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
176 |
by (rule lattice.intro, rule dual_semilattice, rule upper_semilattice.intro, rule dual_order) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
177 |
(unfold_locales, auto) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
178 |
|
21733 | 179 |
lemma inf_sup_absorb: "x \<sqinter> (x \<squnion> y) = x" |
25102
db3e412c4cb1
antisymmetry not a default intro rule any longer
haftmann
parents:
25062
diff
changeset
|
180 |
by (blast intro: antisym inf_le1 inf_greatest sup_ge1) |
21733 | 181 |
|
182 |
lemma sup_inf_absorb: "x \<squnion> (x \<sqinter> y) = x" |
|
25102
db3e412c4cb1
antisymmetry not a default intro rule any longer
haftmann
parents:
25062
diff
changeset
|
183 |
by (blast intro: antisym sup_ge1 sup_least inf_le1) |
21733 | 184 |
|
32064 | 185 |
lemmas inf_sup_aci = inf_aci sup_aci |
21734 | 186 |
|
22454 | 187 |
lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 |
188 |
||
21734 | 189 |
text{* Towards distributivity *} |
21249 | 190 |
|
21734 | 191 |
lemma distrib_sup_le: "x \<squnion> (y \<sqinter> z) \<sqsubseteq> (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
32064 | 192 |
by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) |
21734 | 193 |
|
194 |
lemma distrib_inf_le: "(x \<sqinter> y) \<squnion> (x \<sqinter> z) \<sqsubseteq> x \<sqinter> (y \<squnion> z)" |
|
32064 | 195 |
by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) |
21734 | 196 |
|
197 |
text{* If you have one of them, you have them all. *} |
|
21249 | 198 |
|
21733 | 199 |
lemma distrib_imp1: |
21249 | 200 |
assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
201 |
shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
|
202 |
proof- |
|
203 |
have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb) |
|
34209 | 204 |
also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc) |
21249 | 205 |
also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)" |
206 |
by(simp add:inf_sup_absorb inf_commute) |
|
207 |
also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D) |
|
208 |
finally show ?thesis . |
|
209 |
qed |
|
210 |
||
21733 | 211 |
lemma distrib_imp2: |
21249 | 212 |
assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
213 |
shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
|
214 |
proof- |
|
215 |
have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb) |
|
34209 | 216 |
also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc) |
21249 | 217 |
also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)" |
218 |
by(simp add:sup_inf_absorb sup_commute) |
|
219 |
also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D) |
|
220 |
finally show ?thesis . |
|
221 |
qed |
|
222 |
||
21733 | 223 |
end |
21249 | 224 |
|
32568 | 225 |
subsubsection {* Strict order *} |
226 |
||
227 |
context lower_semilattice |
|
228 |
begin |
|
229 |
||
230 |
lemma less_infI1: |
|
231 |
"a \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x" |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
232 |
by (auto simp add: less_le inf_absorb1 intro: le_infI1) |
32568 | 233 |
|
234 |
lemma less_infI2: |
|
235 |
"b \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x" |
|
32642
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
haftmann
parents:
32568
diff
changeset
|
236 |
by (auto simp add: less_le inf_absorb2 intro: le_infI2) |
32568 | 237 |
|
238 |
end |
|
239 |
||
240 |
context upper_semilattice |
|
241 |
begin |
|
242 |
||
243 |
lemma less_supI1: |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
244 |
"x \<sqsubset> a \<Longrightarrow> x \<sqsubset> a \<squnion> b" |
32568 | 245 |
proof - |
246 |
interpret dual: lower_semilattice "op \<ge>" "op >" sup |
|
247 |
by (fact dual_semilattice) |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
248 |
assume "x \<sqsubset> a" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
249 |
then show "x \<sqsubset> a \<squnion> b" |
32568 | 250 |
by (fact dual.less_infI1) |
251 |
qed |
|
252 |
||
253 |
lemma less_supI2: |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
254 |
"x \<sqsubset> b \<Longrightarrow> x \<sqsubset> a \<squnion> b" |
32568 | 255 |
proof - |
256 |
interpret dual: lower_semilattice "op \<ge>" "op >" sup |
|
257 |
by (fact dual_semilattice) |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
258 |
assume "x \<sqsubset> b" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
259 |
then show "x \<sqsubset> a \<squnion> b" |
32568 | 260 |
by (fact dual.less_infI2) |
261 |
qed |
|
262 |
||
263 |
end |
|
264 |
||
21249 | 265 |
|
24164 | 266 |
subsection {* Distributive lattices *} |
21249 | 267 |
|
22454 | 268 |
class distrib_lattice = lattice + |
21249 | 269 |
assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
270 |
||
21733 | 271 |
context distrib_lattice |
272 |
begin |
|
273 |
||
274 |
lemma sup_inf_distrib2: |
|
21249 | 275 |
"(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" |
32064 | 276 |
by(simp add: inf_sup_aci sup_inf_distrib1) |
21249 | 277 |
|
21733 | 278 |
lemma inf_sup_distrib1: |
21249 | 279 |
"x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
280 |
by(rule distrib_imp2[OF sup_inf_distrib1]) |
|
281 |
||
21733 | 282 |
lemma inf_sup_distrib2: |
21249 | 283 |
"(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)" |
32064 | 284 |
by(simp add: inf_sup_aci inf_sup_distrib1) |
21249 | 285 |
|
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
286 |
lemma dual_distrib_lattice: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
287 |
"distrib_lattice (op \<ge>) (op >) sup inf" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
288 |
by (rule distrib_lattice.intro, rule dual_lattice) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
289 |
(unfold_locales, fact inf_sup_distrib1) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
290 |
|
21733 | 291 |
lemmas distrib = |
21249 | 292 |
sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 |
293 |
||
21733 | 294 |
end |
295 |
||
21249 | 296 |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
297 |
subsection {* Bounded lattices and boolean algebras *} |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
298 |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
299 |
class bounded_lattice = lattice + top + bot |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
300 |
begin |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
301 |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
302 |
lemma dual_bounded_lattice: |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
303 |
"bounded_lattice (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
304 |
by (rule bounded_lattice.intro, rule dual_lattice) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
305 |
(unfold_locales, auto simp add: less_le_not_le) |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
306 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
307 |
lemma inf_bot_left [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
308 |
"\<bottom> \<sqinter> x = \<bottom>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
309 |
by (rule inf_absorb1) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
310 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
311 |
lemma inf_bot_right [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
312 |
"x \<sqinter> \<bottom> = \<bottom>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
313 |
by (rule inf_absorb2) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
314 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
315 |
lemma sup_top_left [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
316 |
"\<top> \<squnion> x = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
317 |
by (rule sup_absorb1) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
318 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
319 |
lemma sup_top_right [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
320 |
"x \<squnion> \<top> = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
321 |
by (rule sup_absorb2) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
322 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
323 |
lemma inf_top_left [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
324 |
"\<top> \<sqinter> x = x" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
325 |
by (rule inf_absorb2) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
326 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
327 |
lemma inf_top_right [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
328 |
"x \<sqinter> \<top> = x" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
329 |
by (rule inf_absorb1) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
330 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
331 |
lemma sup_bot_left [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
332 |
"\<bottom> \<squnion> x = x" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
333 |
by (rule sup_absorb2) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
334 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
335 |
lemma sup_bot_right [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
336 |
"x \<squnion> \<bottom> = x" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
337 |
by (rule sup_absorb1) simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
338 |
|
32568 | 339 |
lemma inf_eq_top_eq1: |
340 |
assumes "A \<sqinter> B = \<top>" |
|
341 |
shows "A = \<top>" |
|
342 |
proof (cases "B = \<top>") |
|
343 |
case True with assms show ?thesis by simp |
|
344 |
next |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
345 |
case False with top_greatest have "B \<sqsubset> \<top>" by (auto intro: neq_le_trans) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
346 |
then have "A \<sqinter> B \<sqsubset> \<top>" by (rule less_infI2) |
32568 | 347 |
with assms show ?thesis by simp |
348 |
qed |
|
349 |
||
350 |
lemma inf_eq_top_eq2: |
|
351 |
assumes "A \<sqinter> B = \<top>" |
|
352 |
shows "B = \<top>" |
|
353 |
by (rule inf_eq_top_eq1, unfold inf_commute [of B]) (fact assms) |
|
354 |
||
355 |
lemma sup_eq_bot_eq1: |
|
356 |
assumes "A \<squnion> B = \<bottom>" |
|
357 |
shows "A = \<bottom>" |
|
358 |
proof - |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
359 |
interpret dual: bounded_lattice "op \<ge>" "op >" "op \<squnion>" "op \<sqinter>" \<top> \<bottom> |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
360 |
by (rule dual_bounded_lattice) |
32568 | 361 |
from dual.inf_eq_top_eq1 assms show ?thesis . |
362 |
qed |
|
363 |
||
364 |
lemma sup_eq_bot_eq2: |
|
365 |
assumes "A \<squnion> B = \<bottom>" |
|
366 |
shows "B = \<bottom>" |
|
367 |
proof - |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
368 |
interpret dual: bounded_lattice "op \<ge>" "op >" "op \<squnion>" "op \<sqinter>" \<top> \<bottom> |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
369 |
by (rule dual_bounded_lattice) |
32568 | 370 |
from dual.inf_eq_top_eq2 assms show ?thesis . |
371 |
qed |
|
372 |
||
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
373 |
end |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
374 |
|
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
375 |
class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
376 |
assumes inf_compl_bot: "x \<sqinter> - x = \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
377 |
and sup_compl_top: "x \<squnion> - x = \<top>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
378 |
assumes diff_eq: "x - y = x \<sqinter> - y" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
379 |
begin |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
380 |
|
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
381 |
lemma dual_boolean_algebra: |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
382 |
"boolean_algebra (\<lambda>x y. x \<squnion> - y) uminus (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
383 |
by (rule boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
384 |
(unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
385 |
|
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
386 |
lemma compl_inf_bot: |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
387 |
"- x \<sqinter> x = \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
388 |
by (simp add: inf_commute inf_compl_bot) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
389 |
|
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
390 |
lemma compl_sup_top: |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
391 |
"- x \<squnion> x = \<top>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
392 |
by (simp add: sup_commute sup_compl_top) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
393 |
|
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
394 |
lemma compl_unique: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
395 |
assumes "x \<sqinter> y = \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
396 |
and "x \<squnion> y = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
397 |
shows "- x = y" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
398 |
proof - |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
399 |
have "(x \<sqinter> - x) \<squnion> (- x \<sqinter> y) = (x \<sqinter> y) \<squnion> (- x \<sqinter> y)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
400 |
using inf_compl_bot assms(1) by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
401 |
then have "(- x \<sqinter> x) \<squnion> (- x \<sqinter> y) = (y \<sqinter> x) \<squnion> (y \<sqinter> - x)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
402 |
by (simp add: inf_commute) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
403 |
then have "- x \<sqinter> (x \<squnion> y) = y \<sqinter> (x \<squnion> - x)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
404 |
by (simp add: inf_sup_distrib1) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
405 |
then have "- x \<sqinter> \<top> = y \<sqinter> \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
406 |
using sup_compl_top assms(2) by simp |
34209 | 407 |
then show "- x = y" by simp |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
408 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
409 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
410 |
lemma double_compl [simp]: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
411 |
"- (- x) = x" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
412 |
using compl_inf_bot compl_sup_top by (rule compl_unique) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
413 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
414 |
lemma compl_eq_compl_iff [simp]: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
415 |
"- x = - y \<longleftrightarrow> x = y" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
416 |
proof |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
417 |
assume "- x = - y" |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
418 |
then have "- x \<sqinter> y = \<bottom>" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
419 |
and "- x \<squnion> y = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
420 |
by (simp_all add: compl_inf_bot compl_sup_top) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
421 |
then have "- (- x) = y" by (rule compl_unique) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
422 |
then show "x = y" by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
423 |
next |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
424 |
assume "x = y" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
425 |
then show "- x = - y" by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
426 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
427 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
428 |
lemma compl_bot_eq [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
429 |
"- \<bottom> = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
430 |
proof - |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
431 |
from sup_compl_top have "\<bottom> \<squnion> - \<bottom> = \<top>" . |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
432 |
then show ?thesis by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
433 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
434 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
435 |
lemma compl_top_eq [simp]: |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
436 |
"- \<top> = \<bottom>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
437 |
proof - |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
438 |
from inf_compl_bot have "\<top> \<sqinter> - \<top> = \<bottom>" . |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
439 |
then show ?thesis by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
440 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
441 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
442 |
lemma compl_inf [simp]: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
443 |
"- (x \<sqinter> y) = - x \<squnion> - y" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
444 |
proof (rule compl_unique) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
445 |
have "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = ((x \<sqinter> y) \<sqinter> - x) \<squnion> ((x \<sqinter> y) \<sqinter> - y)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
446 |
by (rule inf_sup_distrib1) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
447 |
also have "... = (y \<sqinter> (x \<sqinter> - x)) \<squnion> (x \<sqinter> (y \<sqinter> - y))" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
448 |
by (simp only: inf_commute inf_assoc inf_left_commute) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
449 |
finally show "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = \<bottom>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
450 |
by (simp add: inf_compl_bot) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
451 |
next |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
452 |
have "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = (x \<squnion> (- x \<squnion> - y)) \<sqinter> (y \<squnion> (- x \<squnion> - y))" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
453 |
by (rule sup_inf_distrib2) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
454 |
also have "... = (- y \<squnion> (x \<squnion> - x)) \<sqinter> (- x \<squnion> (y \<squnion> - y))" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
455 |
by (simp only: sup_commute sup_assoc sup_left_commute) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
456 |
finally show "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = \<top>" |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
457 |
by (simp add: sup_compl_top) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
458 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
459 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
460 |
lemma compl_sup [simp]: |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
461 |
"- (x \<squnion> y) = - x \<sqinter> - y" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
462 |
proof - |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
463 |
interpret boolean_algebra "\<lambda>x y. x \<squnion> - y" uminus "op \<ge>" "op >" "op \<squnion>" "op \<sqinter>" \<top> \<bottom> |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
464 |
by (rule dual_boolean_algebra) |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
465 |
then show ?thesis by simp |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
466 |
qed |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
467 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
468 |
end |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
469 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
470 |
|
22454 | 471 |
subsection {* Uniqueness of inf and sup *} |
472 |
||
22737 | 473 |
lemma (in lower_semilattice) inf_unique: |
22454 | 474 |
fixes f (infixl "\<triangle>" 70) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
475 |
assumes le1: "\<And>x y. x \<triangle> y \<sqsubseteq> x" and le2: "\<And>x y. x \<triangle> y \<sqsubseteq> y" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
476 |
and greatest: "\<And>x y z. x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<triangle> z" |
22737 | 477 |
shows "x \<sqinter> y = x \<triangle> y" |
22454 | 478 |
proof (rule antisym) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
479 |
show "x \<triangle> y \<sqsubseteq> x \<sqinter> y" by (rule le_infI) (rule le1, rule le2) |
22454 | 480 |
next |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
481 |
have leI: "\<And>x y z. x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<triangle> z" by (blast intro: greatest) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
482 |
show "x \<sqinter> y \<sqsubseteq> x \<triangle> y" by (rule leI) simp_all |
22454 | 483 |
qed |
484 |
||
22737 | 485 |
lemma (in upper_semilattice) sup_unique: |
22454 | 486 |
fixes f (infixl "\<nabla>" 70) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
487 |
assumes ge1 [simp]: "\<And>x y. x \<sqsubseteq> x \<nabla> y" and ge2: "\<And>x y. y \<sqsubseteq> x \<nabla> y" |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
488 |
and least: "\<And>x y z. y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<nabla> z \<sqsubseteq> x" |
22737 | 489 |
shows "x \<squnion> y = x \<nabla> y" |
22454 | 490 |
proof (rule antisym) |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
491 |
show "x \<squnion> y \<sqsubseteq> x \<nabla> y" by (rule le_supI) (rule ge1, rule ge2) |
22454 | 492 |
next |
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
493 |
have leI: "\<And>x y z. x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<nabla> y \<sqsubseteq> z" by (blast intro: least) |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
494 |
show "x \<nabla> y \<sqsubseteq> x \<squnion> y" by (rule leI) simp_all |
22454 | 495 |
qed |
496 |
||
497 |
||
22916 | 498 |
subsection {* @{const min}/@{const max} on linear orders as |
499 |
special case of @{const inf}/@{const sup} *} |
|
500 |
||
32512 | 501 |
sublocale linorder < min_max!: distrib_lattice less_eq less min max |
28823 | 502 |
proof |
22916 | 503 |
fix x y z |
32512 | 504 |
show "max x (min y z) = min (max x y) (max x z)" |
505 |
by (auto simp add: min_def max_def) |
|
22916 | 506 |
qed (auto simp add: min_def max_def not_le less_imp_le) |
21249 | 507 |
|
22454 | 508 |
lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)" |
25102
db3e412c4cb1
antisymmetry not a default intro rule any longer
haftmann
parents:
25062
diff
changeset
|
509 |
by (rule ext)+ (auto intro: antisym) |
21733 | 510 |
|
22454 | 511 |
lemma sup_max: "sup = (max \<Colon> 'a\<Colon>{upper_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)" |
25102
db3e412c4cb1
antisymmetry not a default intro rule any longer
haftmann
parents:
25062
diff
changeset
|
512 |
by (rule ext)+ (auto intro: antisym) |
21733 | 513 |
|
21249 | 514 |
lemmas le_maxI1 = min_max.sup_ge1 |
515 |
lemmas le_maxI2 = min_max.sup_ge2 |
|
21381 | 516 |
|
21249 | 517 |
lemmas max_ac = min_max.sup_assoc min_max.sup_commute |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
518 |
mk_left_commute [of max, OF min_max.sup_assoc min_max.sup_commute] |
21249 | 519 |
|
520 |
lemmas min_ac = min_max.inf_assoc min_max.inf_commute |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
521 |
mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] |
21249 | 522 |
|
22454 | 523 |
|
524 |
subsection {* Bool as lattice *} |
|
525 |
||
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
526 |
instantiation bool :: boolean_algebra |
25510 | 527 |
begin |
528 |
||
529 |
definition |
|
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
530 |
bool_Compl_def: "uminus = Not" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
531 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
532 |
definition |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
533 |
bool_diff_def: "A - B \<longleftrightarrow> A \<and> \<not> B" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
534 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
535 |
definition |
25510 | 536 |
inf_bool_eq: "P \<sqinter> Q \<longleftrightarrow> P \<and> Q" |
537 |
||
538 |
definition |
|
539 |
sup_bool_eq: "P \<squnion> Q \<longleftrightarrow> P \<or> Q" |
|
540 |
||
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
541 |
instance proof |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
542 |
qed (simp_all add: inf_bool_eq sup_bool_eq le_bool_def |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
543 |
bot_bool_eq top_bool_eq bool_Compl_def bool_diff_def, auto) |
22454 | 544 |
|
25510 | 545 |
end |
546 |
||
32781 | 547 |
lemma sup_boolI1: |
548 |
"P \<Longrightarrow> P \<squnion> Q" |
|
549 |
by (simp add: sup_bool_eq) |
|
550 |
||
551 |
lemma sup_boolI2: |
|
552 |
"Q \<Longrightarrow> P \<squnion> Q" |
|
553 |
by (simp add: sup_bool_eq) |
|
554 |
||
555 |
lemma sup_boolE: |
|
556 |
"P \<squnion> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" |
|
557 |
by (auto simp add: sup_bool_eq) |
|
558 |
||
23878 | 559 |
|
560 |
subsection {* Fun as lattice *} |
|
561 |
||
25510 | 562 |
instantiation "fun" :: (type, lattice) lattice |
563 |
begin |
|
564 |
||
565 |
definition |
|
28562 | 566 |
inf_fun_eq [code del]: "f \<sqinter> g = (\<lambda>x. f x \<sqinter> g x)" |
25510 | 567 |
|
568 |
definition |
|
28562 | 569 |
sup_fun_eq [code del]: "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)" |
25510 | 570 |
|
32780 | 571 |
instance proof |
572 |
qed (simp_all add: le_fun_def inf_fun_eq sup_fun_eq) |
|
23878 | 573 |
|
25510 | 574 |
end |
23878 | 575 |
|
576 |
instance "fun" :: (type, distrib_lattice) distrib_lattice |
|
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
577 |
proof |
32780 | 578 |
qed (simp_all add: inf_fun_eq sup_fun_eq sup_inf_distrib1) |
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
579 |
|
34007
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
580 |
instance "fun" :: (type, bounded_lattice) bounded_lattice .. |
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents:
32781
diff
changeset
|
581 |
|
31991
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
582 |
instantiation "fun" :: (type, uminus) uminus |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
583 |
begin |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
584 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
585 |
definition |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
586 |
fun_Compl_def: "- A = (\<lambda>x. - A x)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
587 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
588 |
instance .. |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
589 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
590 |
end |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
591 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
592 |
instantiation "fun" :: (type, minus) minus |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
593 |
begin |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
594 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
595 |
definition |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
596 |
fun_diff_def: "A - B = (\<lambda>x. A x - B x)" |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
597 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
598 |
instance .. |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
599 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
600 |
end |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
601 |
|
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
602 |
instance "fun" :: (type, boolean_algebra) boolean_algebra |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
603 |
proof |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
604 |
qed (simp_all add: inf_fun_eq sup_fun_eq bot_fun_eq top_fun_eq fun_Compl_def fun_diff_def |
37390299214a
added boolean_algebra type class; tuned lattice duals
haftmann
parents:
30729
diff
changeset
|
605 |
inf_compl_bot sup_compl_top diff_eq) |
23878 | 606 |
|
26794 | 607 |
|
25062 | 608 |
no_notation |
25382 | 609 |
less_eq (infix "\<sqsubseteq>" 50) and |
610 |
less (infix "\<sqsubset>" 50) and |
|
611 |
inf (infixl "\<sqinter>" 70) and |
|
32568 | 612 |
sup (infixl "\<squnion>" 65) and |
613 |
top ("\<top>") and |
|
614 |
bot ("\<bottom>") |
|
25062 | 615 |
|
21249 | 616 |
end |