author | ballarin |
Tue, 18 Sep 2007 18:52:17 +0200 | |
changeset 24640 | 85a6c200ecd3 |
parent 24514 | 540eaf87e42d |
child 24749 | 151b3758f576 |
permissions | -rw-r--r-- |
21249 | 1 |
(* Title: HOL/Lattices.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
*) |
|
5 |
||
22454 | 6 |
header {* Abstract lattices *} |
21249 | 7 |
|
8 |
theory Lattices |
|
9 |
imports Orderings |
|
10 |
begin |
|
11 |
||
12 |
subsection{* Lattices *} |
|
13 |
||
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
14 |
class lower_semilattice = order + |
21249 | 15 |
fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70) |
22737 | 16 |
assumes inf_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x" |
17 |
and inf_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y" |
|
21733 | 18 |
and inf_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z" |
21249 | 19 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
20 |
class upper_semilattice = order + |
21249 | 21 |
fixes sup :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) |
22737 | 22 |
assumes sup_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y" |
23 |
and sup_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y" |
|
21733 | 24 |
and sup_least: "y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<squnion> z \<sqsubseteq> x" |
21249 | 25 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
26 |
class lattice = lower_semilattice + upper_semilattice |
21249 | 27 |
|
21733 | 28 |
subsubsection{* Intro and elim rules*} |
29 |
||
30 |
context lower_semilattice |
|
31 |
begin |
|
21249 | 32 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
33 |
lemmas antisym_intro [intro!] = antisym |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
34 |
lemmas (in -) [rule del] = antisym_intro |
21249 | 35 |
|
21734 | 36 |
lemma le_infI1[intro]: "a \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" |
21733 | 37 |
apply(subgoal_tac "a \<sqinter> b \<sqsubseteq> a") |
22384
33a46e6c7f04
prefix of class interpretation not mandatory any longer
haftmann
parents:
22168
diff
changeset
|
38 |
apply(blast intro: order_trans) |
21733 | 39 |
apply simp |
40 |
done |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
41 |
lemmas (in -) [rule del] = le_infI1 |
21249 | 42 |
|
21734 | 43 |
lemma le_infI2[intro]: "b \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" |
21733 | 44 |
apply(subgoal_tac "a \<sqinter> b \<sqsubseteq> b") |
22384
33a46e6c7f04
prefix of class interpretation not mandatory any longer
haftmann
parents:
22168
diff
changeset
|
45 |
apply(blast intro: order_trans) |
21733 | 46 |
apply simp |
47 |
done |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
48 |
lemmas (in -) [rule del] = le_infI2 |
21733 | 49 |
|
21734 | 50 |
lemma le_infI[intro!]: "x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<sqinter> b" |
21733 | 51 |
by(blast intro: inf_greatest) |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
52 |
lemmas (in -) [rule del] = le_infI |
21249 | 53 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
54 |
lemma le_infE [elim!]: "x \<sqsubseteq> a \<sqinter> b \<Longrightarrow> (x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> P) \<Longrightarrow> P" |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
55 |
by (blast intro: order_trans) |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
56 |
lemmas (in -) [rule del] = le_infE |
21249 | 57 |
|
21734 | 58 |
lemma le_inf_iff [simp]: |
21733 | 59 |
"x \<sqsubseteq> y \<sqinter> z = (x \<sqsubseteq> y \<and> x \<sqsubseteq> z)" |
60 |
by blast |
|
61 |
||
21734 | 62 |
lemma le_iff_inf: "(x \<sqsubseteq> y) = (x \<sqinter> y = x)" |
22168 | 63 |
by(blast dest:eq_iff[THEN iffD1]) |
21249 | 64 |
|
21733 | 65 |
end |
66 |
||
23878 | 67 |
lemma mono_inf: "mono f \<Longrightarrow> f (inf A B) \<le> inf (f A) (f B)" |
68 |
by (auto simp add: mono_def) |
|
69 |
||
21733 | 70 |
|
71 |
context upper_semilattice |
|
72 |
begin |
|
21249 | 73 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
74 |
lemmas antisym_intro [intro!] = antisym |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
75 |
lemmas (in -) [rule del] = antisym_intro |
21249 | 76 |
|
21734 | 77 |
lemma le_supI1[intro]: "x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" |
21733 | 78 |
apply(subgoal_tac "a \<sqsubseteq> a \<squnion> b") |
22384
33a46e6c7f04
prefix of class interpretation not mandatory any longer
haftmann
parents:
22168
diff
changeset
|
79 |
apply(blast intro: order_trans) |
21733 | 80 |
apply simp |
81 |
done |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
82 |
lemmas (in -) [rule del] = le_supI1 |
21249 | 83 |
|
21734 | 84 |
lemma le_supI2[intro]: "x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" |
21733 | 85 |
apply(subgoal_tac "b \<sqsubseteq> a \<squnion> b") |
22384
33a46e6c7f04
prefix of class interpretation not mandatory any longer
haftmann
parents:
22168
diff
changeset
|
86 |
apply(blast intro: order_trans) |
21733 | 87 |
apply simp |
88 |
done |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
89 |
lemmas (in -) [rule del] = le_supI2 |
21733 | 90 |
|
21734 | 91 |
lemma le_supI[intro!]: "a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> a \<squnion> b \<sqsubseteq> x" |
21733 | 92 |
by(blast intro: sup_least) |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
93 |
lemmas (in -) [rule del] = le_supI |
21249 | 94 |
|
21734 | 95 |
lemma le_supE[elim!]: "a \<squnion> b \<sqsubseteq> x \<Longrightarrow> (a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> P) \<Longrightarrow> P" |
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
96 |
by (blast intro: order_trans) |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
97 |
lemmas (in -) [rule del] = le_supE |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
22384
diff
changeset
|
98 |
|
21249 | 99 |
|
21734 | 100 |
lemma ge_sup_conv[simp]: |
21733 | 101 |
"x \<squnion> y \<sqsubseteq> z = (x \<sqsubseteq> z \<and> y \<sqsubseteq> z)" |
102 |
by blast |
|
103 |
||
21734 | 104 |
lemma le_iff_sup: "(x \<sqsubseteq> y) = (x \<squnion> y = y)" |
22168 | 105 |
by(blast dest:eq_iff[THEN iffD1]) |
21734 | 106 |
|
21733 | 107 |
end |
108 |
||
23878 | 109 |
lemma mono_sup: "mono f \<Longrightarrow> sup (f A) (f B) \<le> f (sup A B)" |
110 |
by (auto simp add: mono_def) |
|
111 |
||
21733 | 112 |
|
113 |
subsubsection{* Equational laws *} |
|
21249 | 114 |
|
115 |
||
21733 | 116 |
context lower_semilattice |
117 |
begin |
|
118 |
||
119 |
lemma inf_commute: "(x \<sqinter> y) = (y \<sqinter> x)" |
|
120 |
by blast |
|
121 |
||
122 |
lemma inf_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)" |
|
123 |
by blast |
|
124 |
||
125 |
lemma inf_idem[simp]: "x \<sqinter> x = x" |
|
126 |
by blast |
|
127 |
||
128 |
lemma inf_left_idem[simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y" |
|
129 |
by blast |
|
130 |
||
131 |
lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x" |
|
132 |
by blast |
|
133 |
||
134 |
lemma inf_absorb2: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y" |
|
135 |
by blast |
|
136 |
||
137 |
lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)" |
|
138 |
by blast |
|
139 |
||
140 |
lemmas inf_ACI = inf_commute inf_assoc inf_left_commute inf_left_idem |
|
141 |
||
142 |
end |
|
143 |
||
144 |
||
145 |
context upper_semilattice |
|
146 |
begin |
|
21249 | 147 |
|
21733 | 148 |
lemma sup_commute: "(x \<squnion> y) = (y \<squnion> x)" |
149 |
by blast |
|
150 |
||
151 |
lemma sup_assoc: "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)" |
|
152 |
by blast |
|
153 |
||
154 |
lemma sup_idem[simp]: "x \<squnion> x = x" |
|
155 |
by blast |
|
156 |
||
157 |
lemma sup_left_idem[simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y" |
|
158 |
by blast |
|
159 |
||
160 |
lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x" |
|
161 |
by blast |
|
162 |
||
163 |
lemma sup_absorb2: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y" |
|
164 |
by blast |
|
21249 | 165 |
|
21733 | 166 |
lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)" |
167 |
by blast |
|
168 |
||
169 |
lemmas sup_ACI = sup_commute sup_assoc sup_left_commute sup_left_idem |
|
170 |
||
171 |
end |
|
21249 | 172 |
|
21733 | 173 |
context lattice |
174 |
begin |
|
175 |
||
176 |
lemma inf_sup_absorb: "x \<sqinter> (x \<squnion> y) = x" |
|
177 |
by(blast intro: antisym inf_le1 inf_greatest sup_ge1) |
|
178 |
||
179 |
lemma sup_inf_absorb: "x \<squnion> (x \<sqinter> y) = x" |
|
180 |
by(blast intro: antisym sup_ge1 sup_least inf_le1) |
|
181 |
||
21734 | 182 |
lemmas ACI = inf_ACI sup_ACI |
183 |
||
22454 | 184 |
lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 |
185 |
||
21734 | 186 |
text{* Towards distributivity *} |
21249 | 187 |
|
21734 | 188 |
lemma distrib_sup_le: "x \<squnion> (y \<sqinter> z) \<sqsubseteq> (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
189 |
by blast |
|
190 |
||
191 |
lemma distrib_inf_le: "(x \<sqinter> y) \<squnion> (x \<sqinter> z) \<sqsubseteq> x \<sqinter> (y \<squnion> z)" |
|
192 |
by blast |
|
193 |
||
194 |
||
195 |
text{* If you have one of them, you have them all. *} |
|
21249 | 196 |
|
21733 | 197 |
lemma distrib_imp1: |
21249 | 198 |
assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
199 |
shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
|
200 |
proof- |
|
201 |
have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb) |
|
202 |
also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc) |
|
203 |
also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)" |
|
204 |
by(simp add:inf_sup_absorb inf_commute) |
|
205 |
also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D) |
|
206 |
finally show ?thesis . |
|
207 |
qed |
|
208 |
||
21733 | 209 |
lemma distrib_imp2: |
21249 | 210 |
assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
211 |
shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
|
212 |
proof- |
|
213 |
have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb) |
|
214 |
also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc) |
|
215 |
also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)" |
|
216 |
by(simp add:sup_inf_absorb sup_commute) |
|
217 |
also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D) |
|
218 |
finally show ?thesis . |
|
219 |
qed |
|
220 |
||
21734 | 221 |
(* seems unused *) |
222 |
lemma modular_le: "x \<sqsubseteq> z \<Longrightarrow> x \<squnion> (y \<sqinter> z) \<sqsubseteq> (x \<squnion> y) \<sqinter> z" |
|
223 |
by blast |
|
224 |
||
21733 | 225 |
end |
21249 | 226 |
|
227 |
||
24164 | 228 |
subsection {* Distributive lattices *} |
21249 | 229 |
|
22454 | 230 |
class distrib_lattice = lattice + |
21249 | 231 |
assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
232 |
||
21733 | 233 |
context distrib_lattice |
234 |
begin |
|
235 |
||
236 |
lemma sup_inf_distrib2: |
|
21249 | 237 |
"(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" |
238 |
by(simp add:ACI sup_inf_distrib1) |
|
239 |
||
21733 | 240 |
lemma inf_sup_distrib1: |
21249 | 241 |
"x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
242 |
by(rule distrib_imp2[OF sup_inf_distrib1]) |
|
243 |
||
21733 | 244 |
lemma inf_sup_distrib2: |
21249 | 245 |
"(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)" |
246 |
by(simp add:ACI inf_sup_distrib1) |
|
247 |
||
21733 | 248 |
lemmas distrib = |
21249 | 249 |
sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 |
250 |
||
21733 | 251 |
end |
252 |
||
21249 | 253 |
|
22454 | 254 |
subsection {* Uniqueness of inf and sup *} |
255 |
||
22737 | 256 |
lemma (in lower_semilattice) inf_unique: |
22454 | 257 |
fixes f (infixl "\<triangle>" 70) |
22737 | 258 |
assumes le1: "\<And>x y. x \<triangle> y \<^loc>\<le> x" and le2: "\<And>x y. x \<triangle> y \<^loc>\<le> y" |
259 |
and greatest: "\<And>x y z. x \<^loc>\<le> y \<Longrightarrow> x \<^loc>\<le> z \<Longrightarrow> x \<^loc>\<le> y \<triangle> z" |
|
260 |
shows "x \<sqinter> y = x \<triangle> y" |
|
22454 | 261 |
proof (rule antisym) |
23389 | 262 |
show "x \<triangle> y \<^loc>\<le> x \<sqinter> y" by (rule le_infI) (rule le1, rule le2) |
22454 | 263 |
next |
22737 | 264 |
have leI: "\<And>x y z. x \<^loc>\<le> y \<Longrightarrow> x \<^loc>\<le> z \<Longrightarrow> x \<^loc>\<le> y \<triangle> z" by (blast intro: greatest) |
265 |
show "x \<sqinter> y \<^loc>\<le> x \<triangle> y" by (rule leI) simp_all |
|
22454 | 266 |
qed |
267 |
||
22737 | 268 |
lemma (in upper_semilattice) sup_unique: |
22454 | 269 |
fixes f (infixl "\<nabla>" 70) |
22737 | 270 |
assumes ge1 [simp]: "\<And>x y. x \<^loc>\<le> x \<nabla> y" and ge2: "\<And>x y. y \<^loc>\<le> x \<nabla> y" |
271 |
and least: "\<And>x y z. y \<^loc>\<le> x \<Longrightarrow> z \<^loc>\<le> x \<Longrightarrow> y \<nabla> z \<^loc>\<le> x" |
|
272 |
shows "x \<squnion> y = x \<nabla> y" |
|
22454 | 273 |
proof (rule antisym) |
23389 | 274 |
show "x \<squnion> y \<^loc>\<le> x \<nabla> y" by (rule le_supI) (rule ge1, rule ge2) |
22454 | 275 |
next |
22737 | 276 |
have leI: "\<And>x y z. x \<^loc>\<le> z \<Longrightarrow> y \<^loc>\<le> z \<Longrightarrow> x \<nabla> y \<^loc>\<le> z" by (blast intro: least) |
277 |
show "x \<nabla> y \<^loc>\<le> x \<squnion> y" by (rule leI) simp_all |
|
22454 | 278 |
qed |
279 |
||
280 |
||
22916 | 281 |
subsection {* @{const min}/@{const max} on linear orders as |
282 |
special case of @{const inf}/@{const sup} *} |
|
283 |
||
284 |
lemma (in linorder) distrib_lattice_min_max: |
|
23018
1d29bc31b0cb
no special treatment in naming of locale predicates stemming form classes
haftmann
parents:
22916
diff
changeset
|
285 |
"distrib_lattice (op \<^loc>\<le>) (op \<^loc><) min max" |
22916 | 286 |
proof unfold_locales |
287 |
have aux: "\<And>x y \<Colon> 'a. x \<^loc>< y \<Longrightarrow> y \<^loc>\<le> x \<Longrightarrow> x = y" |
|
288 |
by (auto simp add: less_le antisym) |
|
289 |
fix x y z |
|
290 |
show "max x (min y z) = min (max x y) (max x z)" |
|
291 |
unfolding min_def max_def |
|
24640
85a6c200ecd3
Simplified proofs due to transitivity reasoner setup.
ballarin
parents:
24514
diff
changeset
|
292 |
by auto |
22916 | 293 |
qed (auto simp add: min_def max_def not_le less_imp_le) |
21249 | 294 |
|
295 |
interpretation min_max: |
|
22454 | 296 |
distrib_lattice ["op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max] |
23948 | 297 |
by (rule distrib_lattice_min_max) |
21249 | 298 |
|
22454 | 299 |
lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)" |
300 |
by (rule ext)+ auto |
|
21733 | 301 |
|
22454 | 302 |
lemma sup_max: "sup = (max \<Colon> 'a\<Colon>{upper_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)" |
303 |
by (rule ext)+ auto |
|
21733 | 304 |
|
21249 | 305 |
lemmas le_maxI1 = min_max.sup_ge1 |
306 |
lemmas le_maxI2 = min_max.sup_ge2 |
|
21381 | 307 |
|
21249 | 308 |
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
|
309 |
mk_left_commute [of max, OF min_max.sup_assoc min_max.sup_commute] |
21249 | 310 |
|
311 |
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
|
312 |
mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] |
21249 | 313 |
|
22454 | 314 |
text {* |
315 |
Now we have inherited antisymmetry as an intro-rule on all |
|
316 |
linear orders. This is a problem because it applies to bool, which is |
|
317 |
undesirable. |
|
318 |
*} |
|
319 |
||
320 |
lemmas [rule del] = min_max.antisym_intro min_max.le_infI min_max.le_supI |
|
321 |
min_max.le_supE min_max.le_infE min_max.le_supI1 min_max.le_supI2 |
|
322 |
min_max.le_infI1 min_max.le_infI2 |
|
323 |
||
324 |
||
23878 | 325 |
subsection {* Complete lattices *} |
326 |
||
327 |
class complete_lattice = lattice + |
|
328 |
fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900) |
|
24345 | 329 |
and Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900) |
23878 | 330 |
assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x" |
24345 | 331 |
and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A" |
332 |
assumes Sup_upper: "x \<in> A \<Longrightarrow> x \<sqsubseteq> \<Squnion>A" |
|
333 |
and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> \<Squnion>A \<sqsubseteq> z" |
|
23878 | 334 |
begin |
335 |
||
336 |
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<^loc>\<le> a}" |
|
24345 | 337 |
by (auto intro: Inf_lower Inf_greatest Sup_upper Sup_least) |
23878 | 338 |
|
24345 | 339 |
lemma Sup_Inf: "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<^loc>\<le> b}" |
340 |
by (auto intro: Inf_lower Inf_greatest Sup_upper Sup_least) |
|
23878 | 341 |
|
342 |
lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}" |
|
24345 | 343 |
unfolding Sup_Inf by auto |
23878 | 344 |
|
345 |
lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}" |
|
346 |
unfolding Inf_Sup by auto |
|
347 |
||
348 |
lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A" |
|
349 |
apply (rule antisym) |
|
350 |
apply (rule le_infI) |
|
351 |
apply (rule Inf_lower) |
|
352 |
apply simp |
|
353 |
apply (rule Inf_greatest) |
|
354 |
apply (rule Inf_lower) |
|
355 |
apply simp |
|
356 |
apply (rule Inf_greatest) |
|
357 |
apply (erule insertE) |
|
358 |
apply (rule le_infI1) |
|
359 |
apply simp |
|
360 |
apply (rule le_infI2) |
|
361 |
apply (erule Inf_lower) |
|
362 |
done |
|
363 |
||
24345 | 364 |
lemma Sup_insert: "\<Squnion>insert a A = a \<squnion> \<Squnion>A" |
23878 | 365 |
apply (rule antisym) |
366 |
apply (rule Sup_least) |
|
367 |
apply (erule insertE) |
|
368 |
apply (rule le_supI1) |
|
369 |
apply simp |
|
370 |
apply (rule le_supI2) |
|
371 |
apply (erule Sup_upper) |
|
372 |
apply (rule le_supI) |
|
373 |
apply (rule Sup_upper) |
|
374 |
apply simp |
|
375 |
apply (rule Sup_least) |
|
376 |
apply (rule Sup_upper) |
|
377 |
apply simp |
|
378 |
done |
|
379 |
||
380 |
lemma Inf_singleton [simp]: |
|
381 |
"\<Sqinter>{a} = a" |
|
382 |
by (auto intro: antisym Inf_lower Inf_greatest) |
|
383 |
||
24345 | 384 |
lemma Sup_singleton [simp]: |
23878 | 385 |
"\<Squnion>{a} = a" |
386 |
by (auto intro: antisym Sup_upper Sup_least) |
|
387 |
||
388 |
lemma Inf_insert_simp: |
|
389 |
"\<Sqinter>insert a A = (if A = {} then a else a \<sqinter> \<Sqinter>A)" |
|
390 |
by (cases "A = {}") (simp_all, simp add: Inf_insert) |
|
391 |
||
392 |
lemma Sup_insert_simp: |
|
393 |
"\<Squnion>insert a A = (if A = {} then a else a \<squnion> \<Squnion>A)" |
|
394 |
by (cases "A = {}") (simp_all, simp add: Sup_insert) |
|
395 |
||
396 |
lemma Inf_binary: |
|
397 |
"\<Sqinter>{a, b} = a \<sqinter> b" |
|
398 |
by (simp add: Inf_insert_simp) |
|
399 |
||
400 |
lemma Sup_binary: |
|
401 |
"\<Squnion>{a, b} = a \<squnion> b" |
|
402 |
by (simp add: Sup_insert_simp) |
|
403 |
||
404 |
end |
|
405 |
||
406 |
definition |
|
407 |
top :: "'a::complete_lattice" |
|
408 |
where |
|
409 |
"top = Inf {}" |
|
410 |
||
411 |
definition |
|
412 |
bot :: "'a::complete_lattice" |
|
413 |
where |
|
414 |
"bot = Sup {}" |
|
415 |
||
416 |
lemma top_greatest [simp]: "x \<le> top" |
|
417 |
by (unfold top_def, rule Inf_greatest, simp) |
|
418 |
||
419 |
lemma bot_least [simp]: "bot \<le> x" |
|
420 |
by (unfold bot_def, rule Sup_least, simp) |
|
421 |
||
422 |
definition |
|
423 |
SUPR :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" |
|
424 |
where |
|
425 |
"SUPR A f == Sup (f ` A)" |
|
426 |
||
427 |
definition |
|
428 |
INFI :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" |
|
429 |
where |
|
430 |
"INFI A f == Inf (f ` A)" |
|
431 |
||
432 |
syntax |
|
433 |
"_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) |
|
434 |
"_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) |
|
435 |
"_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) |
|
436 |
"_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) |
|
437 |
||
438 |
translations |
|
439 |
"SUP x y. B" == "SUP x. SUP y. B" |
|
440 |
"SUP x. B" == "CONST SUPR UNIV (%x. B)" |
|
441 |
"SUP x. B" == "SUP x:UNIV. B" |
|
442 |
"SUP x:A. B" == "CONST SUPR A (%x. B)" |
|
443 |
"INF x y. B" == "INF x. INF y. B" |
|
444 |
"INF x. B" == "CONST INFI UNIV (%x. B)" |
|
445 |
"INF x. B" == "INF x:UNIV. B" |
|
446 |
"INF x:A. B" == "CONST INFI A (%x. B)" |
|
447 |
||
448 |
(* To avoid eta-contraction of body: *) |
|
449 |
print_translation {* |
|
450 |
let |
|
451 |
fun btr' syn (A :: Abs abs :: ts) = |
|
452 |
let val (x,t) = atomic_abs_tr' abs |
|
453 |
in list_comb (Syntax.const syn $ x $ A $ t, ts) end |
|
454 |
val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const |
|
455 |
in |
|
456 |
[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] |
|
457 |
end |
|
458 |
*} |
|
459 |
||
460 |
lemma le_SUPI: "i : A \<Longrightarrow> M i \<le> (SUP i:A. M i)" |
|
461 |
by (auto simp add: SUPR_def intro: Sup_upper) |
|
462 |
||
463 |
lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<le> u) \<Longrightarrow> (SUP i:A. M i) \<le> u" |
|
464 |
by (auto simp add: SUPR_def intro: Sup_least) |
|
465 |
||
466 |
lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<le> M i" |
|
467 |
by (auto simp add: INFI_def intro: Inf_lower) |
|
468 |
||
469 |
lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<le> M i) \<Longrightarrow> u \<le> (INF i:A. M i)" |
|
470 |
by (auto simp add: INFI_def intro: Inf_greatest) |
|
471 |
||
472 |
lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M" |
|
473 |
by (auto intro: order_antisym SUP_leI le_SUPI) |
|
474 |
||
475 |
lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M" |
|
476 |
by (auto intro: order_antisym INF_leI le_INFI) |
|
477 |
||
478 |
||
22454 | 479 |
subsection {* Bool as lattice *} |
480 |
||
481 |
instance bool :: distrib_lattice |
|
482 |
inf_bool_eq: "inf P Q \<equiv> P \<and> Q" |
|
483 |
sup_bool_eq: "sup P Q \<equiv> P \<or> Q" |
|
484 |
by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def) |
|
485 |
||
23878 | 486 |
instance bool :: complete_lattice |
487 |
Inf_bool_def: "Inf A \<equiv> \<forall>x\<in>A. x" |
|
24345 | 488 |
Sup_bool_def: "Sup A \<equiv> \<exists>x\<in>A. x" |
489 |
by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) |
|
23878 | 490 |
|
491 |
lemma Inf_empty_bool [simp]: |
|
492 |
"Inf {}" |
|
493 |
unfolding Inf_bool_def by auto |
|
494 |
||
495 |
lemma not_Sup_empty_bool [simp]: |
|
496 |
"\<not> Sup {}" |
|
24345 | 497 |
unfolding Sup_bool_def by auto |
23878 | 498 |
|
499 |
lemma top_bool_eq: "top = True" |
|
500 |
by (iprover intro!: order_antisym le_boolI top_greatest) |
|
501 |
||
502 |
lemma bot_bool_eq: "bot = False" |
|
503 |
by (iprover intro!: order_antisym le_boolI bot_least) |
|
504 |
||
505 |
||
506 |
subsection {* Set as lattice *} |
|
507 |
||
508 |
instance set :: (type) distrib_lattice |
|
509 |
inf_set_eq: "inf A B \<equiv> A \<inter> B" |
|
510 |
sup_set_eq: "sup A B \<equiv> A \<union> B" |
|
511 |
by intro_classes (auto simp add: inf_set_eq sup_set_eq) |
|
512 |
||
513 |
lemmas [code func del] = inf_set_eq sup_set_eq |
|
514 |
||
24514
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
515 |
lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B" |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
516 |
apply (fold inf_set_eq sup_set_eq) |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
517 |
apply (erule mono_inf) |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
518 |
done |
23878 | 519 |
|
24514
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
520 |
lemma mono_Un: "mono f \<Longrightarrow> f A \<union> f B \<subseteq> f (A \<union> B)" |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
521 |
apply (fold inf_set_eq sup_set_eq) |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
522 |
apply (erule mono_sup) |
540eaf87e42d
mono_Int/Un: proper proof, avoid illegal schematic type vars;
wenzelm
parents:
24345
diff
changeset
|
523 |
done |
23878 | 524 |
|
525 |
instance set :: (type) complete_lattice |
|
526 |
Inf_set_def: "Inf S \<equiv> \<Inter>S" |
|
24345 | 527 |
Sup_set_def: "Sup S \<equiv> \<Union>S" |
528 |
by intro_classes (auto simp add: Inf_set_def Sup_set_def) |
|
23878 | 529 |
|
24345 | 530 |
lemmas [code func del] = Inf_set_def Sup_set_def |
23878 | 531 |
|
532 |
lemma top_set_eq: "top = UNIV" |
|
533 |
by (iprover intro!: subset_antisym subset_UNIV top_greatest) |
|
534 |
||
535 |
lemma bot_set_eq: "bot = {}" |
|
536 |
by (iprover intro!: subset_antisym empty_subsetI bot_least) |
|
537 |
||
538 |
||
539 |
subsection {* Fun as lattice *} |
|
540 |
||
541 |
instance "fun" :: (type, lattice) lattice |
|
542 |
inf_fun_eq: "inf f g \<equiv> (\<lambda>x. inf (f x) (g x))" |
|
543 |
sup_fun_eq: "sup f g \<equiv> (\<lambda>x. sup (f x) (g x))" |
|
544 |
apply intro_classes |
|
545 |
unfolding inf_fun_eq sup_fun_eq |
|
546 |
apply (auto intro: le_funI) |
|
547 |
apply (rule le_funI) |
|
548 |
apply (auto dest: le_funD) |
|
549 |
apply (rule le_funI) |
|
550 |
apply (auto dest: le_funD) |
|
551 |
done |
|
552 |
||
553 |
lemmas [code func del] = inf_fun_eq sup_fun_eq |
|
554 |
||
555 |
instance "fun" :: (type, distrib_lattice) distrib_lattice |
|
556 |
by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) |
|
557 |
||
558 |
instance "fun" :: (type, complete_lattice) complete_lattice |
|
559 |
Inf_fun_def: "Inf A \<equiv> (\<lambda>x. Inf {y. \<exists>f\<in>A. y = f x})" |
|
24345 | 560 |
Sup_fun_def: "Sup A \<equiv> (\<lambda>x. Sup {y. \<exists>f\<in>A. y = f x})" |
561 |
by intro_classes |
|
562 |
(auto simp add: Inf_fun_def Sup_fun_def le_fun_def |
|
563 |
intro: Inf_lower Sup_upper Inf_greatest Sup_least) |
|
23878 | 564 |
|
24345 | 565 |
lemmas [code func del] = Inf_fun_def Sup_fun_def |
23878 | 566 |
|
567 |
lemma Inf_empty_fun: |
|
568 |
"Inf {} = (\<lambda>_. Inf {})" |
|
569 |
by rule (auto simp add: Inf_fun_def) |
|
570 |
||
571 |
lemma Sup_empty_fun: |
|
572 |
"Sup {} = (\<lambda>_. Sup {})" |
|
24345 | 573 |
by rule (auto simp add: Sup_fun_def) |
23878 | 574 |
|
575 |
lemma top_fun_eq: "top = (\<lambda>x. top)" |
|
576 |
by (iprover intro!: order_antisym le_funI top_greatest) |
|
577 |
||
578 |
lemma bot_fun_eq: "bot = (\<lambda>x. bot)" |
|
579 |
by (iprover intro!: order_antisym le_funI bot_least) |
|
580 |
||
581 |
||
582 |
text {* redundant bindings *} |
|
22454 | 583 |
|
584 |
lemmas inf_aci = inf_ACI |
|
585 |
lemmas sup_aci = sup_ACI |
|
586 |
||
21249 | 587 |
end |