src/HOL/Ord.thy
author paulson
Wed, 09 Oct 1996 13:32:33 +0200
changeset 2073 fb0655539d05
parent 2006 72754e060aa2
child 2259 e6d738f2b9a9
permissions -rw-r--r--
New unified treatment of sequent calculi by Sara Kalvala combines the old LK and Modal with the new ILL (Int. Linear Logic)

(*  Title:      HOL/Ord.thy
    ID:         $Id$
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

The type class for ordered types    (* FIXME improve comment *)
*)

Ord = HOL +

axclass
  ord < term

consts
  "op <"        :: ['a::ord, 'a] => bool             ("(_/ < _)"  [50,51] 50)
  "op <="       :: ['a::ord, 'a] => bool             ("(_/ <= _)" [50,51] 50)
  mono          :: ['a::ord => 'b::ord] => bool       (*monotonicity*)
  min, max      :: ['a::ord, 'a] => 'a

defs
  mono_def      "mono(f) == (!A B. A <= B --> f(A) <= f(B))"
  min_def       "min a b == (if a <= b then a else b)"
  max_def       "max a b == (if a <= b then b else a)"

end