Ord.thy
author lcp
Tue, 06 Sep 1994 10:54:46 +0200
changeset 134 4b7da5a895e7
parent 118 5b96b1252cdc
child 145 a9f7ff3a464c
permissions -rw-r--r--
HOL/ind_syntax/factors: now returns only factors in the product type that associate to the right. Previously the proof of the induction rule crashed on types such as (bool*bool)*bool.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
118
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     1
(*  Title: 	HOL/Ord.thy
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     2
    ID:         $Id$
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     3
    Author: 	Tobias Nipkow, Cambridge University Computer Laboratory
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     4
    Copyright   1993  University of Cambridge
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     5
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     6
The type class for ordered types
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     7
*)
5b96b1252cdc HOL/Ord.thy,.ML: files now have header comments
lcp
parents: 0
diff changeset
     8
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
Ord = HOL +
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    10
classes
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    11
  ord < term
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    12
consts
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    13
  "<", "<="     :: "['a::ord, 'a] => bool"              (infixl 50)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    14
  mono		:: "['a::ord => 'b::ord] => bool"       (*monotonicity*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    15
  min,max	:: "['a::ord,'a] => 'a"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    16
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    17
rules
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    18
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    19
mono_def  "mono(f)  == (!A B. A <= B --> f(A) <= f(B))"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    20
min_def   "min(a,b) == if(a <= b, a, b)"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    21
max_def   "max(a,b) == if(a <= b, b, a)"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    22
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    23
end