Ord.thy
author lcp
Thu, 18 Aug 1994 11:43:40 +0200
changeset 107 960e332d2e70
parent 0 7949f97df77a
child 118 5b96b1252cdc
permissions -rw-r--r--
HOL/Sum: rotated arguments of sum_case; added translation for case macro HOL/Sum: now has Part primitives, moved from ex/Simult, with extra laws from ZF/Sum

Ord = HOL +
classes
  ord < term
consts
  "<", "<="     :: "['a::ord, 'a] => bool"              (infixl 50)
  mono		:: "['a::ord => 'b::ord] => bool"       (*monotonicity*)
  min,max	:: "['a::ord,'a] => 'a"

rules

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

end