| author | wenzelm | 
| Thu, 28 Mar 2024 12:56:20 +0100 | |
| changeset 80051 | 9cdc4f533b91 | 
| parent 77811 | ae9e6218443d | 
| permissions | -rw-r--r-- | 
| 56796 | 1  | 
(* Title: HOL/Library/Finite_Lattice.thy  | 
2  | 
Author: Alessandro Coglio  | 
|
3  | 
*)  | 
|
| 50634 | 4  | 
|
| 77811 | 5  | 
section \<open>Finite Lattices\<close>  | 
6  | 
||
| 50634 | 7  | 
theory Finite_Lattice  | 
| 
51115
 
7dbd6832a689
consolidation of library theories on product orders
 
haftmann 
parents: 
50634 
diff
changeset
 | 
8  | 
imports Product_Order  | 
| 50634 | 9  | 
begin  | 
10  | 
||
| 77811 | 11  | 
subsection \<open>Finite Complete Lattices\<close>  | 
12  | 
||
| 60500 | 13  | 
text \<open>A non-empty finite lattice is a complete lattice.  | 
| 50634 | 14  | 
Since types are never empty in Isabelle/HOL,  | 
| 69593 | 15  | 
a type of classes \<^class>\<open>finite\<close> and \<^class>\<open>lattice\<close>  | 
16  | 
should also have class \<^class>\<open>complete_lattice\<close>.  | 
|
| 50634 | 17  | 
A type class is defined  | 
| 69593 | 18  | 
that extends classes \<^class>\<open>finite\<close> and \<^class>\<open>lattice\<close>  | 
19  | 
with the operators \<^const>\<open>bot\<close>, \<^const>\<open>top\<close>, \<^const>\<open>Inf\<close>, and \<^const>\<open>Sup\<close>,  | 
|
| 50634 | 20  | 
along with assumptions that define these operators  | 
| 69593 | 21  | 
in terms of the ones of classes \<^class>\<open>finite\<close> and \<^class>\<open>lattice\<close>.  | 
22  | 
The resulting class is a subclass of \<^class>\<open>complete_lattice\<close>.\<close>  | 
|
| 50634 | 23  | 
|
24  | 
class finite_lattice_complete = finite + lattice + bot + top + Inf + Sup +  | 
|
| 56796 | 25  | 
assumes bot_def: "bot = Inf_fin UNIV"  | 
26  | 
assumes top_def: "top = Sup_fin UNIV"  | 
|
27  | 
assumes Inf_def: "Inf A = Finite_Set.fold inf top A"  | 
|
28  | 
assumes Sup_def: "Sup A = Finite_Set.fold sup bot A"  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
29  | 
|
| 60500 | 30  | 
text \<open>The definitional assumptions  | 
| 69593 | 31  | 
on the operators \<^const>\<open>bot\<close> and \<^const>\<open>top\<close>  | 
32  | 
of class \<^class>\<open>finite_lattice_complete\<close>  | 
|
| 60500 | 33  | 
ensure that they yield bottom and top.\<close>  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
34  | 
|
| 56796 | 35  | 
lemma finite_lattice_complete_bot_least: "(bot::'a::finite_lattice_complete) \<le> x"  | 
36  | 
by (auto simp: bot_def intro: Inf_fin.coboundedI)  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
37  | 
|
| 
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
38  | 
instance finite_lattice_complete \<subseteq> order_bot  | 
| 60679 | 39  | 
by standard (auto simp: finite_lattice_complete_bot_least)  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
40  | 
|
| 56796 | 41  | 
lemma finite_lattice_complete_top_greatest: "(top::'a::finite_lattice_complete) \<ge> x"  | 
42  | 
by (auto simp: top_def Sup_fin.coboundedI)  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
43  | 
|
| 
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
44  | 
instance finite_lattice_complete \<subseteq> order_top  | 
| 60679 | 45  | 
by standard (auto simp: finite_lattice_complete_top_greatest)  | 
| 50634 | 46  | 
|
47  | 
instance finite_lattice_complete \<subseteq> bounded_lattice ..  | 
|
48  | 
||
| 60500 | 49  | 
text \<open>The definitional assumptions  | 
| 69593 | 50  | 
on the operators \<^const>\<open>Inf\<close> and \<^const>\<open>Sup\<close>  | 
51  | 
of class \<^class>\<open>finite_lattice_complete\<close>  | 
|
| 60500 | 52  | 
ensure that they yield infimum and supremum.\<close>  | 
| 50634 | 53  | 
|
| 56796 | 54  | 
lemma finite_lattice_complete_Inf_empty: "Inf {} = (top :: 'a::finite_lattice_complete)"
 | 
| 51489 | 55  | 
by (simp add: Inf_def)  | 
56  | 
||
| 56796 | 57  | 
lemma finite_lattice_complete_Sup_empty: "Sup {} = (bot :: 'a::finite_lattice_complete)"
 | 
| 51489 | 58  | 
by (simp add: Sup_def)  | 
59  | 
||
60  | 
lemma finite_lattice_complete_Inf_insert:  | 
|
61  | 
fixes A :: "'a::finite_lattice_complete set"  | 
|
62  | 
shows "Inf (insert x A) = inf x (Inf A)"  | 
|
63  | 
proof -  | 
|
| 56796 | 64  | 
interpret comp_fun_idem "inf :: 'a \<Rightarrow> _"  | 
65  | 
by (fact comp_fun_idem_inf)  | 
|
| 51489 | 66  | 
show ?thesis by (simp add: Inf_def)  | 
67  | 
qed  | 
|
68  | 
||
69  | 
lemma finite_lattice_complete_Sup_insert:  | 
|
70  | 
fixes A :: "'a::finite_lattice_complete set"  | 
|
71  | 
shows "Sup (insert x A) = sup x (Sup A)"  | 
|
72  | 
proof -  | 
|
| 56796 | 73  | 
interpret comp_fun_idem "sup :: 'a \<Rightarrow> _"  | 
74  | 
by (fact comp_fun_idem_sup)  | 
|
| 51489 | 75  | 
show ?thesis by (simp add: Sup_def)  | 
76  | 
qed  | 
|
77  | 
||
| 50634 | 78  | 
lemma finite_lattice_complete_Inf_lower:  | 
79  | 
"(x::'a::finite_lattice_complete) \<in> A \<Longrightarrow> Inf A \<le> x"  | 
|
| 56796 | 80  | 
using finite [of A]  | 
81  | 
by (induct A) (auto simp add: finite_lattice_complete_Inf_insert intro: le_infI2)  | 
|
| 50634 | 82  | 
|
83  | 
lemma finite_lattice_complete_Inf_greatest:  | 
|
84  | 
"\<forall>x::'a::finite_lattice_complete \<in> A. z \<le> x \<Longrightarrow> z \<le> Inf A"  | 
|
| 56796 | 85  | 
using finite [of A]  | 
86  | 
by (induct A) (auto simp add: finite_lattice_complete_Inf_empty finite_lattice_complete_Inf_insert)  | 
|
| 50634 | 87  | 
|
88  | 
lemma finite_lattice_complete_Sup_upper:  | 
|
89  | 
"(x::'a::finite_lattice_complete) \<in> A \<Longrightarrow> Sup A \<ge> x"  | 
|
| 56796 | 90  | 
using finite [of A]  | 
91  | 
by (induct A) (auto simp add: finite_lattice_complete_Sup_insert intro: le_supI2)  | 
|
| 50634 | 92  | 
|
93  | 
lemma finite_lattice_complete_Sup_least:  | 
|
94  | 
"\<forall>x::'a::finite_lattice_complete \<in> A. z \<ge> x \<Longrightarrow> z \<ge> Sup A"  | 
|
| 56796 | 95  | 
using finite [of A]  | 
96  | 
by (induct A) (auto simp add: finite_lattice_complete_Sup_empty finite_lattice_complete_Sup_insert)  | 
|
| 50634 | 97  | 
|
98  | 
instance finite_lattice_complete \<subseteq> complete_lattice  | 
|
99  | 
proof  | 
|
100  | 
qed (auto simp:  | 
|
| 56796 | 101  | 
finite_lattice_complete_Inf_lower  | 
102  | 
finite_lattice_complete_Inf_greatest  | 
|
103  | 
finite_lattice_complete_Sup_upper  | 
|
104  | 
finite_lattice_complete_Sup_least  | 
|
105  | 
finite_lattice_complete_Inf_empty  | 
|
106  | 
finite_lattice_complete_Sup_empty)  | 
|
| 50634 | 107  | 
|
| 60500 | 108  | 
text \<open>The product of two finite lattices is already a finite lattice.\<close>  | 
| 50634 | 109  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
110  | 
lemma finite_bot_prod:  | 
| 
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
111  | 
  "(bot :: ('a::finite_lattice_complete \<times> 'b::finite_lattice_complete)) =
 | 
| 56796 | 112  | 
Inf_fin UNIV"  | 
113  | 
by (metis Inf_fin.coboundedI UNIV_I bot.extremum_uniqueI finite_UNIV)  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
114  | 
|
| 
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
115  | 
lemma finite_top_prod:  | 
| 
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
116  | 
  "(top :: ('a::finite_lattice_complete \<times> 'b::finite_lattice_complete)) =
 | 
| 56796 | 117  | 
Sup_fin UNIV"  | 
118  | 
by (metis Sup_fin.coboundedI UNIV_I top.extremum_uniqueI finite_UNIV)  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
119  | 
|
| 50634 | 120  | 
lemma finite_Inf_prod:  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
121  | 
  "Inf(A :: ('a::finite_lattice_complete \<times> 'b::finite_lattice_complete) set) =
 | 
| 56796 | 122  | 
Finite_Set.fold inf top A"  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
123  | 
by (metis Inf_fold_inf finite)  | 
| 50634 | 124  | 
|
125  | 
lemma finite_Sup_prod:  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
126  | 
  "Sup (A :: ('a::finite_lattice_complete \<times> 'b::finite_lattice_complete) set) =
 | 
| 56796 | 127  | 
Finite_Set.fold sup bot A"  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
128  | 
by (metis Sup_fold_sup finite)  | 
| 50634 | 129  | 
|
| 56796 | 130  | 
instance prod :: (finite_lattice_complete, finite_lattice_complete) finite_lattice_complete  | 
| 60679 | 131  | 
by standard (auto simp: finite_bot_prod finite_top_prod finite_Inf_prod finite_Sup_prod)  | 
| 50634 | 132  | 
|
| 60500 | 133  | 
text \<open>Functions with a finite domain and with a finite lattice as codomain  | 
134  | 
already form a finite lattice.\<close>  | 
|
| 50634 | 135  | 
|
| 56796 | 136  | 
lemma finite_bot_fun: "(bot :: ('a::finite \<Rightarrow> 'b::finite_lattice_complete)) = Inf_fin UNIV"
 | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
137  | 
by (metis Inf_UNIV Inf_fin_Inf empty_not_UNIV finite)  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
138  | 
|
| 56796 | 139  | 
lemma finite_top_fun: "(top :: ('a::finite \<Rightarrow> 'b::finite_lattice_complete)) = Sup_fin UNIV"
 | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
140  | 
by (metis Sup_UNIV Sup_fin_Sup empty_not_UNIV finite)  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
141  | 
|
| 50634 | 142  | 
lemma finite_Inf_fun:  | 
143  | 
  "Inf (A::('a::finite \<Rightarrow> 'b::finite_lattice_complete) set) =
 | 
|
| 56796 | 144  | 
Finite_Set.fold inf top A"  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
145  | 
by (metis Inf_fold_inf finite)  | 
| 50634 | 146  | 
|
147  | 
lemma finite_Sup_fun:  | 
|
148  | 
  "Sup (A::('a::finite \<Rightarrow> 'b::finite_lattice_complete) set) =
 | 
|
| 56796 | 149  | 
Finite_Set.fold sup bot A"  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
150  | 
by (metis Sup_fold_sup finite)  | 
| 50634 | 151  | 
|
152  | 
instance "fun" :: (finite, finite_lattice_complete) finite_lattice_complete  | 
|
| 60679 | 153  | 
by standard (auto simp: finite_bot_fun finite_top_fun finite_Inf_fun finite_Sup_fun)  | 
| 50634 | 154  | 
|
155  | 
||
| 60500 | 156  | 
subsection \<open>Finite Distributive Lattices\<close>  | 
| 50634 | 157  | 
|
| 60500 | 158  | 
text \<open>A finite distributive lattice is a complete lattice  | 
| 69593 | 159  | 
whose \<^const>\<open>inf\<close> and \<^const>\<open>sup\<close> operators  | 
160  | 
distribute over \<^const>\<open>Sup\<close> and \<^const>\<open>Inf\<close>.\<close>  | 
|
| 50634 | 161  | 
|
162  | 
class finite_distrib_lattice_complete =  | 
|
163  | 
distrib_lattice + finite_lattice_complete  | 
|
164  | 
||
165  | 
lemma finite_distrib_lattice_complete_sup_Inf:  | 
|
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
67829 
diff
changeset
 | 
166  | 
"sup (x::'a::finite_distrib_lattice_complete) (Inf A) = (INF y\<in>A. sup x y)"  | 
| 56796 | 167  | 
using finite  | 
168  | 
by (induct A rule: finite_induct) (simp_all add: sup_inf_distrib1)  | 
|
| 50634 | 169  | 
|
170  | 
lemma finite_distrib_lattice_complete_inf_Sup:  | 
|
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
67829 
diff
changeset
 | 
171  | 
"inf (x::'a::finite_distrib_lattice_complete) (Sup A) = (SUP y\<in>A. inf x y)"  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
60679 
diff
changeset
 | 
172  | 
using finite [of A] by induct (simp_all add: inf_sup_distrib1)  | 
| 50634 | 173  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
174  | 
context finite_distrib_lattice_complete  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
175  | 
begin  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
176  | 
subclass finite_distrib_lattice  | 
| 73832 | 177  | 
proof -  | 
178  | 
show "class.finite_distrib_lattice Inf Sup inf (\<le>) (<) sup bot top"  | 
|
179  | 
proof  | 
|
180  | 
show "bot = Inf UNIV"  | 
|
181  | 
unfolding bot_def top_def Inf_def  | 
|
182  | 
using Inf_fin.eq_fold Inf_fin.insert inf.absorb2 by force  | 
|
183  | 
next  | 
|
184  | 
show "top = Sup UNIV"  | 
|
185  | 
unfolding bot_def top_def Sup_def  | 
|
186  | 
using Sup_fin.eq_fold Sup_fin.insert by force  | 
|
187  | 
next  | 
|
188  | 
    show "Inf {} = Sup UNIV"
 | 
|
189  | 
unfolding Inf_def Sup_def bot_def top_def  | 
|
190  | 
using Sup_fin.eq_fold Sup_fin.insert by force  | 
|
191  | 
next  | 
|
192  | 
    show "Sup {} = Inf UNIV"
 | 
|
193  | 
unfolding Inf_def Sup_def bot_def top_def  | 
|
194  | 
using Inf_fin.eq_fold Inf_fin.insert inf.absorb2 by force  | 
|
195  | 
next  | 
|
196  | 
interpret comp_fun_idem_inf: comp_fun_idem inf  | 
|
197  | 
by (fact comp_fun_idem_inf)  | 
|
198  | 
show "Inf (insert a A) = inf a (Inf A)" for a A  | 
|
199  | 
using comp_fun_idem_inf.fold_insert_idem Inf_def finite by simp  | 
|
200  | 
next  | 
|
201  | 
interpret comp_fun_idem_sup: comp_fun_idem sup  | 
|
202  | 
by (fact comp_fun_idem_sup)  | 
|
203  | 
show "Sup (insert a A) = sup a (Sup A)" for a A  | 
|
204  | 
using comp_fun_idem_sup.fold_insert_idem Sup_def finite by simp  | 
|
205  | 
qed  | 
|
206  | 
qed  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
207  | 
end  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
208  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62343 
diff
changeset
 | 
209  | 
instance finite_distrib_lattice_complete \<subseteq> complete_distrib_lattice ..  | 
| 50634 | 210  | 
|
| 60500 | 211  | 
text \<open>The product of two finite distributive lattices  | 
212  | 
is already a finite distributive lattice.\<close>  | 
|
| 50634 | 213  | 
|
214  | 
instance prod ::  | 
|
215  | 
(finite_distrib_lattice_complete, finite_distrib_lattice_complete)  | 
|
216  | 
finite_distrib_lattice_complete  | 
|
| 56796 | 217  | 
..  | 
| 50634 | 218  | 
|
| 60500 | 219  | 
text \<open>Functions with a finite domain  | 
| 50634 | 220  | 
and with a finite distributive lattice as codomain  | 
| 60500 | 221  | 
already form a finite distributive lattice.\<close>  | 
| 50634 | 222  | 
|
223  | 
instance "fun" ::  | 
|
224  | 
(finite, finite_distrib_lattice_complete) finite_distrib_lattice_complete  | 
|
| 56796 | 225  | 
..  | 
| 50634 | 226  | 
|
| 60500 | 227  | 
subsection \<open>Linear Orders\<close>  | 
| 50634 | 228  | 
|
| 60500 | 229  | 
text \<open>A linear order is a distributive lattice.  | 
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
230  | 
A type class is defined  | 
| 69593 | 231  | 
that extends class \<^class>\<open>linorder\<close>  | 
232  | 
with the operators \<^const>\<open>inf\<close> and \<^const>\<open>sup\<close>,  | 
|
| 50634 | 233  | 
along with assumptions that define these operators  | 
| 69593 | 234  | 
in terms of the ones of class \<^class>\<open>linorder\<close>.  | 
235  | 
The resulting class is a subclass of \<^class>\<open>distrib_lattice\<close>.\<close>  | 
|
| 50634 | 236  | 
|
237  | 
class linorder_lattice = linorder + inf + sup +  | 
|
| 56796 | 238  | 
assumes inf_def: "inf x y = (if x \<le> y then x else y)"  | 
239  | 
assumes sup_def: "sup x y = (if x \<ge> y then x else y)"  | 
|
| 50634 | 240  | 
|
| 60500 | 241  | 
text \<open>The definitional assumptions  | 
| 69593 | 242  | 
on the operators \<^const>\<open>inf\<close> and \<^const>\<open>sup\<close>  | 
243  | 
of class \<^class>\<open>linorder_lattice\<close>  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
244  | 
ensure that they yield infimum and supremum  | 
| 60500 | 245  | 
and that they distribute over each other.\<close>  | 
| 50634 | 246  | 
|
247  | 
lemma linorder_lattice_inf_le1: "inf (x::'a::linorder_lattice) y \<le> x"  | 
|
| 56796 | 248  | 
unfolding inf_def by (metis (full_types) linorder_linear)  | 
| 50634 | 249  | 
|
250  | 
lemma linorder_lattice_inf_le2: "inf (x::'a::linorder_lattice) y \<le> y"  | 
|
| 56796 | 251  | 
unfolding inf_def by (metis (full_types) linorder_linear)  | 
| 50634 | 252  | 
|
253  | 
lemma linorder_lattice_inf_greatest:  | 
|
254  | 
"(x::'a::linorder_lattice) \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> inf y z"  | 
|
| 56796 | 255  | 
unfolding inf_def by (metis (full_types))  | 
| 50634 | 256  | 
|
257  | 
lemma linorder_lattice_sup_ge1: "sup (x::'a::linorder_lattice) y \<ge> x"  | 
|
| 56796 | 258  | 
unfolding sup_def by (metis (full_types) linorder_linear)  | 
| 50634 | 259  | 
|
260  | 
lemma linorder_lattice_sup_ge2: "sup (x::'a::linorder_lattice) y \<ge> y"  | 
|
| 56796 | 261  | 
unfolding sup_def by (metis (full_types) linorder_linear)  | 
| 50634 | 262  | 
|
263  | 
lemma linorder_lattice_sup_least:  | 
|
264  | 
"(x::'a::linorder_lattice) \<ge> y \<Longrightarrow> x \<ge> z \<Longrightarrow> x \<ge> sup y z"  | 
|
| 56796 | 265  | 
by (auto simp: sup_def)  | 
| 50634 | 266  | 
|
267  | 
lemma linorder_lattice_sup_inf_distrib1:  | 
|
268  | 
"sup (x::'a::linorder_lattice) (inf y z) = inf (sup x y) (sup x z)"  | 
|
| 56796 | 269  | 
by (auto simp: inf_def sup_def)  | 
270  | 
||
| 50634 | 271  | 
instance linorder_lattice \<subseteq> distrib_lattice  | 
| 56796 | 272  | 
proof  | 
| 50634 | 273  | 
qed (auto simp:  | 
| 56796 | 274  | 
linorder_lattice_inf_le1  | 
275  | 
linorder_lattice_inf_le2  | 
|
276  | 
linorder_lattice_inf_greatest  | 
|
277  | 
linorder_lattice_sup_ge1  | 
|
278  | 
linorder_lattice_sup_ge2  | 
|
279  | 
linorder_lattice_sup_least  | 
|
280  | 
linorder_lattice_sup_inf_distrib1)  | 
|
| 50634 | 281  | 
|
282  | 
||
| 60500 | 283  | 
subsection \<open>Finite Linear Orders\<close>  | 
| 50634 | 284  | 
|
| 60500 | 285  | 
text \<open>A (non-empty) finite linear order is a complete linear order.\<close>  | 
| 50634 | 286  | 
|
287  | 
class finite_linorder_complete = linorder_lattice + finite_lattice_complete  | 
|
288  | 
||
289  | 
instance finite_linorder_complete \<subseteq> complete_linorder ..  | 
|
290  | 
||
| 60500 | 291  | 
text \<open>A (non-empty) finite linear order is a complete lattice  | 
| 69593 | 292  | 
whose \<^const>\<open>inf\<close> and \<^const>\<open>sup\<close> operators  | 
293  | 
distribute over \<^const>\<open>Sup\<close> and \<^const>\<open>Inf\<close>.\<close>  | 
|
| 50634 | 294  | 
|
295  | 
instance finite_linorder_complete \<subseteq> finite_distrib_lattice_complete ..  | 
|
296  | 
||
297  | 
end  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
51489 
diff
changeset
 | 
298  |