src/HOL/Ord.ML
author nipkow
Thu, 16 Nov 1995 19:50:40 +0100
changeset 1334 32a9fde85699
parent 923 ff1574a81019
child 1465 5d7a7e439cec
permissions -rw-r--r--
added rev_contrapos
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.ML
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
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
open Ord;
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
val [prem] = goalw Ord.thy [mono_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
    "[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
by (REPEAT (ares_tac [allI, impI, prem] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
qed "monoI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
val [major,minor] = goalw Ord.thy [mono_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
    "[| mono(f);  A <= B |] ==> f(A) <= f(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
by (rtac (major RS spec RS spec RS mp) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
by (rtac minor 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
qed "monoD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21