src/HOL/Ord.thy
author paulson
Mon, 23 Sep 1996 17:47:49 +0200
changeset 2006 72754e060aa2
parent 1370 7361ac9b024d
child 2259 e6d738f2b9a9
permissions -rw-r--r--
New infix syntax: breaks line BEFORE operator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/Ord.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
The type class for ordered types    (* FIXME improve comment *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
Ord = HOL +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
axclass
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
  ord < term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
consts
2006
72754e060aa2 New infix syntax: breaks line BEFORE operator
paulson
parents: 1370
diff changeset
    15
  "op <"        :: ['a::ord, 'a] => bool             ("(_/ < _)"  [50,51] 50)
72754e060aa2 New infix syntax: breaks line BEFORE operator
paulson
parents: 1370
diff changeset
    16
  "op <="       :: ['a::ord, 'a] => bool             ("(_/ <= _)" [50,51] 50)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 965
diff changeset
    17
  mono          :: ['a::ord => 'b::ord] => bool       (*monotonicity*)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 965
diff changeset
    18
  min, max      :: ['a::ord, 'a] => 'a
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
  mono_def      "mono(f) == (!A B. A <= B --> f(A) <= f(B))"
965
24eef3860714 changed syntax of "if"
clasohm
parents: 923
diff changeset
    22
  min_def       "min a b == (if a <= b then a else b)"
24eef3860714 changed syntax of "if"
clasohm
parents: 923
diff changeset
    23
  max_def       "max a b == (if a <= b then b else a)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26