src/HOL/AxClasses/Lattice/LatInsts.thy
author paulson
Tue, 16 Jul 1996 15:49:46 +0200
changeset 1868 836950047d85
parent 1440 de6f18da81bb
child 2606 27cdd600a3b1
permissions -rw-r--r--
Put in minimal simpset to avoid excessive simplification, just as in revision 1.9 of HOL/indrule.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     1
(*  Title:      LatInsts.thy
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     4
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     5
Some lattice instantiations.
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     6
*)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     7
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     8
LatInsts = LatPreInsts +
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     9
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    10
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    11
(* linear orders are lattices *)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    12
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    13
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    14
  lin_order < lattice                   (allI, exI, min_is_inf, max_is_sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    15
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    16
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    17
(* complete lattices are lattices *)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    18
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    19
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    20
  clattice < lattice                    (allI, exI, Inf_is_inf, Sup_is_sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    21
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    22
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    23
(* products of lattices are lattices *)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    24
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    25
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    26
  "*" :: (lattice, lattice) lattice     (allI, exI, prod_is_inf, prod_is_sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    27
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    28
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    29
  fun :: (term, lattice) lattice        (allI, exI, fun_is_inf, fun_is_sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    30
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    31
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    32
(* duals of lattices are lattices *)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    33
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    34
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    35
  dual :: (lattice) lattice             (allI, exI, dual_is_inf, dual_is_sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    36
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    37
(*FIXME bug workaround (see also OrdInsts.thy)*)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    38
instance
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    39
  dual :: (lin_order) lin_order         (le_dual_lin)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    40
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    41
end