src/HOL/AxClasses/Lattice/CLattice.thy
author wenzelm
Tue, 27 May 1997 15:45:07 +0200
changeset 3362 0b268cff9344
parent 2606 27cdd600a3b1
permissions -rw-r--r--
NJ 1.09.2x as factory default!
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     1
(*  Title:      CLattice.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
Complete lattices are orders with infima and suprema of arbitrary
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     6
subsets.
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     7
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     8
TODO:
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
     9
  derive some more well-known theorems (e.g. ex_Inf == ex_Sup)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    10
*)
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    11
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    12
CLattice = Order +
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    13
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    14
axclass
2606
27cdd600a3b1 tuned names: partial order, linear order;
wenzelm
parents: 1573
diff changeset
    15
  clattice < partial_order
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    16
  ex_Inf       "ALL A. EX inf. is_Inf A inf"
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    17
  ex_Sup       "ALL A. EX sup. is_Sup A sup"
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    18
1573
6d66b59f94a9 added constdefs section
clasohm
parents: 1440
diff changeset
    19
constdefs
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    20
  Inf           :: "'a::clattice set => 'a"
1573
6d66b59f94a9 added constdefs section
clasohm
parents: 1440
diff changeset
    21
  "Inf A == @inf. is_Inf A inf"
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    22
1573
6d66b59f94a9 added constdefs section
clasohm
parents: 1440
diff changeset
    23
  Sup           :: "'a::clattice set => 'a"
6d66b59f94a9 added constdefs section
clasohm
parents: 1440
diff changeset
    24
  "Sup A == @sup. is_Sup A sup"
1440
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    25
de6f18da81bb added this stuff;
wenzelm
parents:
diff changeset
    26
end