|
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
|
|
|
9 |
imports Set
|
|
|
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 |
|
|
|
55 |
lemma (in lattice) inf_sup_absorb: "x \<sqinter> (x \<squnion> y) = x"
|
|
|
56 |
by(blast intro: antisym inf_le1 inf_least sup_ge1)
|
|
|
57 |
|
|
|
58 |
lemma (in lattice) sup_inf_absorb: "x \<squnion> (x \<sqinter> y) = x"
|
|
|
59 |
by(blast intro: antisym sup_ge1 sup_greatest inf_le1)
|
|
|
60 |
|
|
|
61 |
lemma (in lower_semilattice) inf_absorb: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
|
|
|
62 |
by(blast intro: antisym inf_le1 inf_least refl)
|
|
|
63 |
|
|
|
64 |
lemma (in upper_semilattice) sup_absorb: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
|
|
|
65 |
by(blast intro: antisym sup_ge2 sup_greatest refl)
|
|
|
66 |
|
|
|
67 |
text{* Towards distributivity: if you have one of them, you have them all. *}
|
|
|
68 |
|
|
|
69 |
lemma (in lattice) distrib_imp1:
|
|
|
70 |
assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
|
|
|
71 |
shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
|
|
|
72 |
proof-
|
|
|
73 |
have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb)
|
|
|
74 |
also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc)
|
|
|
75 |
also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)"
|
|
|
76 |
by(simp add:inf_sup_absorb inf_commute)
|
|
|
77 |
also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D)
|
|
|
78 |
finally show ?thesis .
|
|
|
79 |
qed
|
|
|
80 |
|
|
|
81 |
lemma (in lattice) distrib_imp2:
|
|
|
82 |
assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
|
|
|
83 |
shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
|
|
|
84 |
proof-
|
|
|
85 |
have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb)
|
|
|
86 |
also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc)
|
|
|
87 |
also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)"
|
|
|
88 |
by(simp add:sup_inf_absorb sup_commute)
|
|
|
89 |
also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D)
|
|
|
90 |
finally show ?thesis .
|
|
|
91 |
qed
|
|
|
92 |
|
|
|
93 |
text{* A package of rewrite rules for deciding equivalence wrt ACI: *}
|
|
|
94 |
|
|
|
95 |
lemma (in lower_semilattice) inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)"
|
|
|
96 |
proof -
|
|
|
97 |
have "x \<sqinter> (y \<sqinter> z) = (y \<sqinter> z) \<sqinter> x" by (simp only: inf_commute)
|
|
|
98 |
also have "... = y \<sqinter> (z \<sqinter> x)" by (simp only: inf_assoc)
|
|
|
99 |
also have "z \<sqinter> x = x \<sqinter> z" by (simp only: inf_commute)
|
|
|
100 |
finally show ?thesis .
|
|
|
101 |
qed
|
|
|
102 |
|
|
|
103 |
lemma (in upper_semilattice) sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
|
|
|
104 |
proof -
|
|
|
105 |
have "x \<squnion> (y \<squnion> z) = (y \<squnion> z) \<squnion> x" by (simp only: sup_commute)
|
|
|
106 |
also have "... = y \<squnion> (z \<squnion> x)" by (simp only: sup_assoc)
|
|
|
107 |
also have "z \<squnion> x = x \<squnion> z" by (simp only: sup_commute)
|
|
|
108 |
finally show ?thesis .
|
|
|
109 |
qed
|
|
|
110 |
|
|
|
111 |
lemma (in lower_semilattice) inf_left_idem: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
|
|
|
112 |
proof -
|
|
|
113 |
have "x \<sqinter> (x \<sqinter> y) = (x \<sqinter> x) \<sqinter> y" by(simp only:inf_assoc)
|
|
|
114 |
also have "\<dots> = x \<sqinter> y" by(simp)
|
|
|
115 |
finally show ?thesis .
|
|
|
116 |
qed
|
|
|
117 |
|
|
|
118 |
lemma (in upper_semilattice) sup_left_idem: "x \<squnion> (x \<squnion> y) = x \<squnion> y"
|
|
|
119 |
proof -
|
|
|
120 |
have "x \<squnion> (x \<squnion> y) = (x \<squnion> x) \<squnion> y" by(simp only:sup_assoc)
|
|
|
121 |
also have "\<dots> = x \<squnion> y" by(simp)
|
|
|
122 |
finally show ?thesis .
|
|
|
123 |
qed
|
|
|
124 |
|
|
|
125 |
|
|
|
126 |
lemmas (in lower_semilattice) inf_ACI =
|
|
|
127 |
inf_commute inf_assoc inf_left_commute inf_left_idem
|
|
|
128 |
|
|
|
129 |
lemmas (in upper_semilattice) sup_ACI =
|
|
|
130 |
sup_commute sup_assoc sup_left_commute sup_left_idem
|
|
|
131 |
|
|
|
132 |
lemmas (in lattice) ACI = inf_ACI sup_ACI
|
|
|
133 |
|
|
|
134 |
|
|
|
135 |
subsection{* Distributive lattices *}
|
|
|
136 |
|
|
|
137 |
locale distrib_lattice = lattice +
|
|
|
138 |
assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
|
|
|
139 |
|
|
|
140 |
lemma (in distrib_lattice) sup_inf_distrib2:
|
|
|
141 |
"(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)"
|
|
|
142 |
by(simp add:ACI sup_inf_distrib1)
|
|
|
143 |
|
|
|
144 |
lemma (in distrib_lattice) inf_sup_distrib1:
|
|
|
145 |
"x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
|
|
|
146 |
by(rule distrib_imp2[OF sup_inf_distrib1])
|
|
|
147 |
|
|
|
148 |
lemma (in distrib_lattice) inf_sup_distrib2:
|
|
|
149 |
"(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)"
|
|
|
150 |
by(simp add:ACI inf_sup_distrib1)
|
|
|
151 |
|
|
|
152 |
lemmas (in distrib_lattice) distrib =
|
|
|
153 |
sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2
|
|
|
154 |
|
|
|
155 |
|
|
|
156 |
end
|