1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/lattices/".
17 include "ordered_sets.ma".
19 record is_lattice (C:Type) (join,meet:C→C→C) : Prop \def
20 { (* abelian semigroup properties *)
21 l_comm_j: symmetric ? join;
22 l_associative_j: associative ? join;
23 l_comm_m: symmetric ? meet;
24 l_associative_m: associative ? meet;
25 (* other properties *)
26 l_adsorb_j_m: ∀f,g:C. join f (meet f g) = f;
27 l_adsorb_m_j: ∀f,g:C. meet f (join f g) = f
30 record lattice : Type \def
32 l_join: l_carrier→l_carrier→l_carrier;
33 l_meet: l_carrier→l_carrier→l_carrier;
34 l_lattice_properties:> is_lattice ? l_join l_meet
37 interpretation "Lattice meet" 'and a b =
38 (cic:/matita/lattices/l_meet.con _ a b).
40 interpretation "Lattice join" 'or a b =
41 (cic:/matita/lattices/l_join.con _ a b).
43 definition le \def λL:lattice.λf,g:L. (f ∧ g) = f.
45 definition ordered_set_of_lattice: lattice → ordered_set.
50 | apply mk_is_order_relation;
54 rewrite < (l_adsorb_j_m ? ? ? L ? x) in ⊢ (? ? (? ? ? %) ?);
55 rewrite > l_adsorb_m_j;
57 | apply (l_lattice_properties L)
64 rewrite > (l_associative_m ? ? ? L);
67 | unfold antisimmetric;
71 rewrite > (l_comm_m ? ? ? L);
77 coercion cic:/matita/lattices/ordered_set_of_lattice.con.