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