| 1440 |      1 | (*  Title:      LatInsts.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Some lattice instantiations.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | LatInsts = LatPreInsts +
 | 
|  |      9 | 
 | 
|  |     10 | 
 | 
|  |     11 | (* linear orders are lattices *)
 | 
|  |     12 | 
 | 
|  |     13 | instance
 | 
|  |     14 |   lin_order < lattice                   (allI, exI, min_is_inf, max_is_sup)
 | 
|  |     15 | 
 | 
|  |     16 | 
 | 
|  |     17 | (* complete lattices are lattices *)
 | 
|  |     18 | 
 | 
|  |     19 | instance
 | 
|  |     20 |   clattice < lattice                    (allI, exI, Inf_is_inf, Sup_is_sup)
 | 
|  |     21 | 
 | 
|  |     22 | 
 | 
|  |     23 | (* products of lattices are lattices *)
 | 
|  |     24 | 
 | 
|  |     25 | instance
 | 
|  |     26 |   "*" :: (lattice, lattice) lattice     (allI, exI, prod_is_inf, prod_is_sup)
 | 
|  |     27 | 
 | 
|  |     28 | instance
 | 
|  |     29 |   fun :: (term, lattice) lattice        (allI, exI, fun_is_inf, fun_is_sup)
 | 
|  |     30 | 
 | 
|  |     31 | 
 | 
|  |     32 | (* duals of lattices are lattices *)
 | 
|  |     33 | 
 | 
|  |     34 | instance
 | 
|  |     35 |   dual :: (lattice) lattice             (allI, exI, dual_is_inf, dual_is_sup)
 | 
|  |     36 | 
 | 
|  |     37 | (*FIXME bug workaround (see also OrdInsts.thy)*)
 | 
|  |     38 | instance
 | 
|  |     39 |   dual :: (lin_order) lin_order         (le_dual_lin)
 | 
|  |     40 | 
 | 
|  |     41 | end
 |