| author | wenzelm |
| Tue, 31 Jul 2007 13:30:35 +0200 | |
| changeset 24085 | cbad32e7ab40 |
| parent 23948 | 261bd4678076 |
| child 24164 | 911b46812018 |
| 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 |
||
228 |
subsection{* Distributive lattices *}
|
|
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 |
|
292 |
by (auto simp add: intro: antisym, unfold not_le, |
|
293 |
auto intro: less_trans le_less_trans aux) |
|
294 |
qed (auto simp add: min_def max_def not_le less_imp_le) |
|
| 21249 | 295 |
|
296 |
interpretation min_max: |
|
| 22454 | 297 |
distrib_lattice ["op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max] |
| 23948 | 298 |
by (rule distrib_lattice_min_max) |
| 21249 | 299 |
|
| 22454 | 300 |
lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
|
301 |
by (rule ext)+ auto |
|
| 21733 | 302 |
|
| 22454 | 303 |
lemma sup_max: "sup = (max \<Colon> 'a\<Colon>{upper_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
|
304 |
by (rule ext)+ auto |
|
| 21733 | 305 |
|
| 21249 | 306 |
lemmas le_maxI1 = min_max.sup_ge1 |
307 |
lemmas le_maxI2 = min_max.sup_ge2 |
|
| 21381 | 308 |
|
| 21249 | 309 |
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
|
310 |
mk_left_commute [of max, OF min_max.sup_assoc min_max.sup_commute] |
| 21249 | 311 |
|
312 |
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
|
313 |
mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] |
| 21249 | 314 |
|
| 22454 | 315 |
text {*
|
316 |
Now we have inherited antisymmetry as an intro-rule on all |
|
317 |
linear orders. This is a problem because it applies to bool, which is |
|
318 |
undesirable. |
|
319 |
*} |
|
320 |
||
321 |
lemmas [rule del] = min_max.antisym_intro min_max.le_infI min_max.le_supI |
|
322 |
min_max.le_supE min_max.le_infE min_max.le_supI1 min_max.le_supI2 |
|
323 |
min_max.le_infI1 min_max.le_infI2 |
|
324 |
||
325 |
||
| 23878 | 326 |
subsection {* Complete lattices *}
|
327 |
||
328 |
class complete_lattice = lattice + |
|
329 |
fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
|
|
330 |
assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x" |
|
331 |
assumes Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A" |
|
332 |
begin |
|
333 |
||
334 |
definition |
|
335 |
Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
|
|
336 |
where |
|
337 |
"\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<^loc>\<le> b}"
|
|
338 |
||
339 |
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<^loc>\<le> a}"
|
|
340 |
unfolding Sup_def by (auto intro: Inf_greatest Inf_lower) |
|
341 |
||
342 |
lemma Sup_upper: "x \<in> A \<Longrightarrow> x \<^loc>\<le> \<Squnion>A" |
|
343 |
by (auto simp: Sup_def intro: Inf_greatest) |
|
344 |
||
345 |
lemma Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<^loc>\<le> z) \<Longrightarrow> \<Squnion>A \<^loc>\<le> z" |
|
346 |
by (auto simp: Sup_def intro: Inf_lower) |
|
347 |
||
348 |
lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}"
|
|
349 |
unfolding Sup_def by auto |
|
350 |
||
351 |
lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}"
|
|
352 |
unfolding Inf_Sup by auto |
|
353 |
||
354 |
lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A" |
|
355 |
apply (rule antisym) |
|
356 |
apply (rule le_infI) |
|
357 |
apply (rule Inf_lower) |
|
358 |
apply simp |
|
359 |
apply (rule Inf_greatest) |
|
360 |
apply (rule Inf_lower) |
|
361 |
apply simp |
|
362 |
apply (rule Inf_greatest) |
|
363 |
apply (erule insertE) |
|
364 |
apply (rule le_infI1) |
|
365 |
apply simp |
|
366 |
apply (rule le_infI2) |
|
367 |
apply (erule Inf_lower) |
|
368 |
done |
|
369 |
||
| 23948 | 370 |
lemma Sup_insert [code func]: "\<Squnion>insert a A = a \<squnion> \<Squnion>A" |
| 23878 | 371 |
apply (rule antisym) |
372 |
apply (rule Sup_least) |
|
373 |
apply (erule insertE) |
|
374 |
apply (rule le_supI1) |
|
375 |
apply simp |
|
376 |
apply (rule le_supI2) |
|
377 |
apply (erule Sup_upper) |
|
378 |
apply (rule le_supI) |
|
379 |
apply (rule Sup_upper) |
|
380 |
apply simp |
|
381 |
apply (rule Sup_least) |
|
382 |
apply (rule Sup_upper) |
|
383 |
apply simp |
|
384 |
done |
|
385 |
||
386 |
lemma Inf_singleton [simp]: |
|
387 |
"\<Sqinter>{a} = a"
|
|
388 |
by (auto intro: antisym Inf_lower Inf_greatest) |
|
389 |
||
| 23948 | 390 |
lemma Sup_singleton [simp, code func]: |
| 23878 | 391 |
"\<Squnion>{a} = a"
|
392 |
by (auto intro: antisym Sup_upper Sup_least) |
|
393 |
||
394 |
lemma Inf_insert_simp: |
|
395 |
"\<Sqinter>insert a A = (if A = {} then a else a \<sqinter> \<Sqinter>A)"
|
|
396 |
by (cases "A = {}") (simp_all, simp add: Inf_insert)
|
|
397 |
||
398 |
lemma Sup_insert_simp: |
|
399 |
"\<Squnion>insert a A = (if A = {} then a else a \<squnion> \<Squnion>A)"
|
|
400 |
by (cases "A = {}") (simp_all, simp add: Sup_insert)
|
|
401 |
||
402 |
lemma Inf_binary: |
|
403 |
"\<Sqinter>{a, b} = a \<sqinter> b"
|
|
404 |
by (simp add: Inf_insert_simp) |
|
405 |
||
406 |
lemma Sup_binary: |
|
407 |
"\<Squnion>{a, b} = a \<squnion> b"
|
|
408 |
by (simp add: Sup_insert_simp) |
|
409 |
||
410 |
end |
|
411 |
||
412 |
definition |
|
413 |
top :: "'a::complete_lattice" |
|
414 |
where |
|
415 |
"top = Inf {}"
|
|
416 |
||
417 |
definition |
|
418 |
bot :: "'a::complete_lattice" |
|
419 |
where |
|
420 |
"bot = Sup {}"
|
|
421 |
||
422 |
lemma top_greatest [simp]: "x \<le> top" |
|
423 |
by (unfold top_def, rule Inf_greatest, simp) |
|
424 |
||
425 |
lemma bot_least [simp]: "bot \<le> x" |
|
426 |
by (unfold bot_def, rule Sup_least, simp) |
|
427 |
||
428 |
definition |
|
429 |
SUPR :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b"
|
|
430 |
where |
|
431 |
"SUPR A f == Sup (f ` A)" |
|
432 |
||
433 |
definition |
|
434 |
INFI :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b"
|
|
435 |
where |
|
436 |
"INFI A f == Inf (f ` A)" |
|
437 |
||
438 |
syntax |
|
439 |
"_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10)
|
|
440 |
"_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10)
|
|
441 |
"_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10)
|
|
442 |
"_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10)
|
|
443 |
||
444 |
translations |
|
445 |
"SUP x y. B" == "SUP x. SUP y. B" |
|
446 |
"SUP x. B" == "CONST SUPR UNIV (%x. B)" |
|
447 |
"SUP x. B" == "SUP x:UNIV. B" |
|
448 |
"SUP x:A. B" == "CONST SUPR A (%x. B)" |
|
449 |
"INF x y. B" == "INF x. INF y. B" |
|
450 |
"INF x. B" == "CONST INFI UNIV (%x. B)" |
|
451 |
"INF x. B" == "INF x:UNIV. B" |
|
452 |
"INF x:A. B" == "CONST INFI A (%x. B)" |
|
453 |
||
454 |
(* To avoid eta-contraction of body: *) |
|
455 |
print_translation {*
|
|
456 |
let |
|
457 |
fun btr' syn (A :: Abs abs :: ts) = |
|
458 |
let val (x,t) = atomic_abs_tr' abs |
|
459 |
in list_comb (Syntax.const syn $ x $ A $ t, ts) end |
|
460 |
val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const
|
|
461 |
in |
|
462 |
[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")]
|
|
463 |
end |
|
464 |
*} |
|
465 |
||
466 |
lemma le_SUPI: "i : A \<Longrightarrow> M i \<le> (SUP i:A. M i)" |
|
467 |
by (auto simp add: SUPR_def intro: Sup_upper) |
|
468 |
||
469 |
lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<le> u) \<Longrightarrow> (SUP i:A. M i) \<le> u" |
|
470 |
by (auto simp add: SUPR_def intro: Sup_least) |
|
471 |
||
472 |
lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<le> M i" |
|
473 |
by (auto simp add: INFI_def intro: Inf_lower) |
|
474 |
||
475 |
lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<le> M i) \<Longrightarrow> u \<le> (INF i:A. M i)" |
|
476 |
by (auto simp add: INFI_def intro: Inf_greatest) |
|
477 |
||
478 |
lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
|
|
479 |
by (auto intro: order_antisym SUP_leI le_SUPI) |
|
480 |
||
481 |
lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
|
|
482 |
by (auto intro: order_antisym INF_leI le_INFI) |
|
483 |
||
484 |
||
| 22454 | 485 |
subsection {* Bool as lattice *}
|
486 |
||
487 |
instance bool :: distrib_lattice |
|
488 |
inf_bool_eq: "inf P Q \<equiv> P \<and> Q" |
|
489 |
sup_bool_eq: "sup P Q \<equiv> P \<or> Q" |
|
490 |
by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def) |
|
491 |
||
| 23878 | 492 |
instance bool :: complete_lattice |
493 |
Inf_bool_def: "Inf A \<equiv> \<forall>x\<in>A. x" |
|
494 |
apply intro_classes |
|
495 |
apply (unfold Inf_bool_def) |
|
496 |
apply (iprover intro!: le_boolI elim: ballE) |
|
497 |
apply (iprover intro!: ballI le_boolI elim: ballE le_boolE) |
|
498 |
done |
|
| 22454 | 499 |
|
| 23878 | 500 |
theorem Sup_bool_eq: "Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)" |
501 |
apply (rule order_antisym) |
|
502 |
apply (rule Sup_least) |
|
503 |
apply (rule le_boolI) |
|
504 |
apply (erule bexI, assumption) |
|
505 |
apply (rule le_boolI) |
|
506 |
apply (erule bexE) |
|
507 |
apply (rule le_boolE) |
|
508 |
apply (rule Sup_upper) |
|
509 |
apply assumption+ |
|
510 |
done |
|
511 |
||
512 |
lemma Inf_empty_bool [simp]: |
|
513 |
"Inf {}"
|
|
514 |
unfolding Inf_bool_def by auto |
|
515 |
||
516 |
lemma not_Sup_empty_bool [simp]: |
|
517 |
"\<not> Sup {}"
|
|
518 |
unfolding Sup_def Inf_bool_def by auto |
|
519 |
||
520 |
lemma top_bool_eq: "top = True" |
|
521 |
by (iprover intro!: order_antisym le_boolI top_greatest) |
|
522 |
||
523 |
lemma bot_bool_eq: "bot = False" |
|
524 |
by (iprover intro!: order_antisym le_boolI bot_least) |
|
525 |
||
526 |
||
527 |
subsection {* Set as lattice *}
|
|
528 |
||
529 |
instance set :: (type) distrib_lattice |
|
530 |
inf_set_eq: "inf A B \<equiv> A \<inter> B" |
|
531 |
sup_set_eq: "sup A B \<equiv> A \<union> B" |
|
532 |
by intro_classes (auto simp add: inf_set_eq sup_set_eq) |
|
533 |
||
534 |
lemmas [code func del] = inf_set_eq sup_set_eq |
|
535 |
||
536 |
lemmas mono_Int = mono_inf |
|
537 |
[where ?'a="?'a set", where ?'b="?'b set", unfolded inf_set_eq sup_set_eq] |
|
538 |
||
539 |
lemmas mono_Un = mono_sup |
|
540 |
[where ?'a="?'a set", where ?'b="?'b set", unfolded inf_set_eq sup_set_eq] |
|
541 |
||
542 |
instance set :: (type) complete_lattice |
|
543 |
Inf_set_def: "Inf S \<equiv> \<Inter>S" |
|
544 |
by intro_classes (auto simp add: Inf_set_def) |
|
545 |
||
546 |
lemmas [code func del] = Inf_set_def |
|
547 |
||
548 |
theorem Sup_set_eq: "Sup S = \<Union>S" |
|
549 |
apply (rule subset_antisym) |
|
550 |
apply (rule Sup_least) |
|
551 |
apply (erule Union_upper) |
|
552 |
apply (rule Union_least) |
|
553 |
apply (erule Sup_upper) |
|
554 |
done |
|
555 |
||
556 |
lemma top_set_eq: "top = UNIV" |
|
557 |
by (iprover intro!: subset_antisym subset_UNIV top_greatest) |
|
558 |
||
559 |
lemma bot_set_eq: "bot = {}"
|
|
560 |
by (iprover intro!: subset_antisym empty_subsetI bot_least) |
|
561 |
||
562 |
||
563 |
subsection {* Fun as lattice *}
|
|
564 |
||
565 |
instance "fun" :: (type, lattice) lattice |
|
566 |
inf_fun_eq: "inf f g \<equiv> (\<lambda>x. inf (f x) (g x))" |
|
567 |
sup_fun_eq: "sup f g \<equiv> (\<lambda>x. sup (f x) (g x))" |
|
568 |
apply intro_classes |
|
569 |
unfolding inf_fun_eq sup_fun_eq |
|
570 |
apply (auto intro: le_funI) |
|
571 |
apply (rule le_funI) |
|
572 |
apply (auto dest: le_funD) |
|
573 |
apply (rule le_funI) |
|
574 |
apply (auto dest: le_funD) |
|
575 |
done |
|
576 |
||
577 |
lemmas [code func del] = inf_fun_eq sup_fun_eq |
|
578 |
||
579 |
instance "fun" :: (type, distrib_lattice) distrib_lattice |
|
580 |
by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) |
|
581 |
||
582 |
instance "fun" :: (type, complete_lattice) complete_lattice |
|
583 |
Inf_fun_def: "Inf A \<equiv> (\<lambda>x. Inf {y. \<exists>f\<in>A. y = f x})"
|
|
584 |
apply intro_classes |
|
585 |
apply (unfold Inf_fun_def) |
|
586 |
apply (rule le_funI) |
|
587 |
apply (rule Inf_lower) |
|
588 |
apply (rule CollectI) |
|
589 |
apply (rule bexI) |
|
590 |
apply (rule refl) |
|
591 |
apply assumption |
|
592 |
apply (rule le_funI) |
|
593 |
apply (rule Inf_greatest) |
|
594 |
apply (erule CollectE) |
|
595 |
apply (erule bexE) |
|
596 |
apply (iprover elim: le_funE) |
|
597 |
done |
|
598 |
||
599 |
lemmas [code func del] = Inf_fun_def |
|
600 |
||
601 |
theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y. \<exists>f\<in>A. y = f x})"
|
|
602 |
apply (rule order_antisym) |
|
603 |
apply (rule Sup_least) |
|
604 |
apply (rule le_funI) |
|
605 |
apply (rule Sup_upper) |
|
606 |
apply fast |
|
607 |
apply (rule le_funI) |
|
608 |
apply (rule Sup_least) |
|
609 |
apply (erule CollectE) |
|
610 |
apply (erule bexE) |
|
611 |
apply (drule le_funD [OF Sup_upper]) |
|
612 |
apply simp |
|
613 |
done |
|
614 |
||
615 |
lemma Inf_empty_fun: |
|
616 |
"Inf {} = (\<lambda>_. Inf {})"
|
|
617 |
by rule (auto simp add: Inf_fun_def) |
|
618 |
||
619 |
lemma Sup_empty_fun: |
|
620 |
"Sup {} = (\<lambda>_. Sup {})"
|
|
621 |
proof - |
|
622 |
have aux: "\<And>x. {y. \<exists>f. y = f x} = UNIV" by auto
|
|
623 |
show ?thesis |
|
624 |
by (auto simp add: Sup_def Inf_fun_def Inf_binary inf_bool_eq aux) |
|
625 |
qed |
|
626 |
||
627 |
lemma top_fun_eq: "top = (\<lambda>x. top)" |
|
628 |
by (iprover intro!: order_antisym le_funI top_greatest) |
|
629 |
||
630 |
lemma bot_fun_eq: "bot = (\<lambda>x. bot)" |
|
631 |
by (iprover intro!: order_antisym le_funI bot_least) |
|
632 |
||
633 |
||
634 |
text {* redundant bindings *}
|
|
| 22454 | 635 |
|
636 |
lemmas inf_aci = inf_ACI |
|
637 |
lemmas sup_aci = sup_ACI |
|
638 |
||
| 23878 | 639 |
ML {*
|
640 |
val sup_fun_eq = @{thm sup_fun_eq}
|
|
641 |
val sup_bool_eq = @{thm sup_bool_eq}
|
|
642 |
*} |
|
643 |
||
| 21249 | 644 |
end |