author | wenzelm |
Thu, 22 Dec 2005 00:28:36 +0100 | |
changeset 18457 | 356a9f711899 |
parent 15791 | 446ec11266be |
child 21216 | 1c8580913738 |
permissions | -rw-r--r-- |
15511 | 1 |
(* Title: HOL/Lattices.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
*) |
|
5 |
||
6 |
header {* Lattices via Locales *} |
|
7 |
||
8 |
theory Lattice_Locales |
|
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
9 |
imports HOL |
15511 | 10 |
begin |
11 |
||
12 |
subsection{* Lattices *} |
|
13 |
||
14 |
text{* This theory of lattice locales only defines binary sup and inf |
|
15 |
operations. The extension to finite sets is done in theory @{text |
|
16 |
Finite_Set}. In the longer term it may be better to define arbitrary |
|
17 |
sups and infs via @{text THE}. *} |
|
18 |
||
19 |
locale partial_order = |
|
20 |
fixes below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubseteq>" 50) |
|
21 |
assumes refl[iff]: "x \<sqsubseteq> x" |
|
22 |
and trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" |
|
23 |
and antisym: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" |
|
24 |
||
25 |
locale lower_semilattice = partial_order + |
|
26 |
fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70) |
|
27 |
assumes inf_le1: "x \<sqinter> y \<sqsubseteq> x" and inf_le2: "x \<sqinter> y \<sqsubseteq> y" |
|
28 |
and inf_least: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z" |
|
29 |
||
30 |
locale upper_semilattice = partial_order + |
|
31 |
fixes sup :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) |
|
32 |
assumes sup_ge1: "x \<sqsubseteq> x \<squnion> y" and sup_ge2: "y \<sqsubseteq> x \<squnion> y" |
|
33 |
and sup_greatest: "y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<squnion> z \<sqsubseteq> x" |
|
34 |
||
35 |
locale lattice = lower_semilattice + upper_semilattice |
|
36 |
||
37 |
lemma (in lower_semilattice) inf_commute: "(x \<sqinter> y) = (y \<sqinter> x)" |
|
38 |
by(blast intro: antisym inf_le1 inf_le2 inf_least) |
|
39 |
||
40 |
lemma (in upper_semilattice) sup_commute: "(x \<squnion> y) = (y \<squnion> x)" |
|
41 |
by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest) |
|
42 |
||
43 |
lemma (in lower_semilattice) inf_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)" |
|
44 |
by(blast intro: antisym inf_le1 inf_le2 inf_least trans del:refl) |
|
45 |
||
46 |
lemma (in upper_semilattice) sup_assoc: "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)" |
|
47 |
by(blast intro!: antisym sup_ge1 sup_ge2 intro: sup_greatest trans del:refl) |
|
48 |
||
49 |
lemma (in lower_semilattice) inf_idem[simp]: "x \<sqinter> x = x" |
|
50 |
by(blast intro: antisym inf_le1 inf_le2 inf_least refl) |
|
51 |
||
52 |
lemma (in upper_semilattice) sup_idem[simp]: "x \<squnion> x = x" |
|
53 |
by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest refl) |
|
54 |
||
15791 | 55 |
lemma (in lower_semilattice) inf_left_idem[simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y" |
56 |
by (simp add: inf_assoc[symmetric]) |
|
57 |
||
58 |
lemma (in upper_semilattice) sup_left_idem[simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y" |
|
59 |
by (simp add: sup_assoc[symmetric]) |
|
60 |
||
15511 | 61 |
lemma (in lattice) inf_sup_absorb: "x \<sqinter> (x \<squnion> y) = x" |
62 |
by(blast intro: antisym inf_le1 inf_least sup_ge1) |
|
63 |
||
64 |
lemma (in lattice) sup_inf_absorb: "x \<squnion> (x \<sqinter> y) = x" |
|
65 |
by(blast intro: antisym sup_ge1 sup_greatest inf_le1) |
|
66 |
||
67 |
lemma (in lower_semilattice) inf_absorb: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x" |
|
68 |
by(blast intro: antisym inf_le1 inf_least refl) |
|
69 |
||
70 |
lemma (in upper_semilattice) sup_absorb: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y" |
|
71 |
by(blast intro: antisym sup_ge2 sup_greatest refl) |
|
72 |
||
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
73 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
74 |
lemma (in lower_semilattice) below_inf_conv[simp]: |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
75 |
"x \<sqsubseteq> y \<sqinter> z = (x \<sqsubseteq> y \<and> x \<sqsubseteq> z)" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
76 |
by(blast intro: antisym inf_le1 inf_le2 inf_least refl trans) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
77 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
78 |
lemma (in upper_semilattice) above_sup_conv[simp]: |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
79 |
"x \<squnion> y \<sqsubseteq> z = (x \<sqsubseteq> z \<and> y \<sqsubseteq> z)" |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
80 |
by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest refl trans) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
81 |
|
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
82 |
|
15511 | 83 |
text{* Towards distributivity: if you have one of them, you have them all. *} |
84 |
||
85 |
lemma (in lattice) distrib_imp1: |
|
86 |
assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
|
87 |
shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
|
88 |
proof- |
|
89 |
have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb) |
|
90 |
also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc) |
|
91 |
also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)" |
|
92 |
by(simp add:inf_sup_absorb inf_commute) |
|
93 |
also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D) |
|
94 |
finally show ?thesis . |
|
95 |
qed |
|
96 |
||
97 |
lemma (in lattice) distrib_imp2: |
|
98 |
assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
|
99 |
shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
|
100 |
proof- |
|
101 |
have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb) |
|
102 |
also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc) |
|
103 |
also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)" |
|
104 |
by(simp add:sup_inf_absorb sup_commute) |
|
105 |
also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D) |
|
106 |
finally show ?thesis . |
|
107 |
qed |
|
108 |
||
109 |
text{* A package of rewrite rules for deciding equivalence wrt ACI: *} |
|
110 |
||
111 |
lemma (in lower_semilattice) inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)" |
|
112 |
proof - |
|
113 |
have "x \<sqinter> (y \<sqinter> z) = (y \<sqinter> z) \<sqinter> x" by (simp only: inf_commute) |
|
114 |
also have "... = y \<sqinter> (z \<sqinter> x)" by (simp only: inf_assoc) |
|
115 |
also have "z \<sqinter> x = x \<sqinter> z" by (simp only: inf_commute) |
|
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
116 |
finally(back_subst) show ?thesis . |
15511 | 117 |
qed |
118 |
||
119 |
lemma (in upper_semilattice) sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)" |
|
120 |
proof - |
|
121 |
have "x \<squnion> (y \<squnion> z) = (y \<squnion> z) \<squnion> x" by (simp only: sup_commute) |
|
122 |
also have "... = y \<squnion> (z \<squnion> x)" by (simp only: sup_assoc) |
|
123 |
also have "z \<squnion> x = x \<squnion> z" by (simp only: sup_commute) |
|
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15511
diff
changeset
|
124 |
finally(back_subst) show ?thesis . |
15511 | 125 |
qed |
126 |
||
127 |
lemma (in lower_semilattice) inf_left_idem: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y" |
|
128 |
proof - |
|
129 |
have "x \<sqinter> (x \<sqinter> y) = (x \<sqinter> x) \<sqinter> y" by(simp only:inf_assoc) |
|
130 |
also have "\<dots> = x \<sqinter> y" by(simp) |
|
131 |
finally show ?thesis . |
|
132 |
qed |
|
133 |
||
134 |
lemma (in upper_semilattice) sup_left_idem: "x \<squnion> (x \<squnion> y) = x \<squnion> y" |
|
135 |
proof - |
|
136 |
have "x \<squnion> (x \<squnion> y) = (x \<squnion> x) \<squnion> y" by(simp only:sup_assoc) |
|
137 |
also have "\<dots> = x \<squnion> y" by(simp) |
|
138 |
finally show ?thesis . |
|
139 |
qed |
|
140 |
||
141 |
||
142 |
lemmas (in lower_semilattice) inf_ACI = |
|
143 |
inf_commute inf_assoc inf_left_commute inf_left_idem |
|
144 |
||
145 |
lemmas (in upper_semilattice) sup_ACI = |
|
146 |
sup_commute sup_assoc sup_left_commute sup_left_idem |
|
147 |
||
148 |
lemmas (in lattice) ACI = inf_ACI sup_ACI |
|
149 |
||
150 |
||
151 |
subsection{* Distributive lattices *} |
|
152 |
||
153 |
locale distrib_lattice = lattice + |
|
154 |
assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" |
|
155 |
||
156 |
lemma (in distrib_lattice) sup_inf_distrib2: |
|
157 |
"(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" |
|
158 |
by(simp add:ACI sup_inf_distrib1) |
|
159 |
||
160 |
lemma (in distrib_lattice) inf_sup_distrib1: |
|
161 |
"x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" |
|
162 |
by(rule distrib_imp2[OF sup_inf_distrib1]) |
|
163 |
||
164 |
lemma (in distrib_lattice) inf_sup_distrib2: |
|
165 |
"(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)" |
|
166 |
by(simp add:ACI inf_sup_distrib1) |
|
167 |
||
168 |
lemmas (in distrib_lattice) distrib = |
|
169 |
sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 |
|
170 |
||
171 |
||
172 |
end |