src/HOLCF/Discrete0.ML
author slotosch
Sun, 25 May 1997 11:07:52 +0200
changeset 3323 194ae2e0c193
parent 3018 e65b60b28341
child 5068 fb28eaa07e01
permissions -rw-r--r--
eliminated the constant less by the introduction of the axclass sq_ord added explicit type ::'a::po in the following theorems: minimal2UU,antisym_less_inverse,box_less,not_less2not_eq,monofun_pair and dist_eqI (in domain-package) added instances instance fun :: (term,sq_ord)sq_ord instance "->" :: (term,sq_ord)sq_ord instance "**" :: (sq_ord,sq_ord)sq_ord instance "*" :: (sq_ord,sq_ord)sq_ord instance "++" :: (pcpo,pcpo)sq_ord instance u :: (sq_ord)sq_ord instance lift :: (term)sq_ord instance discr :: (term)sq_ord

(*  Title:      HOLCF/Discrete0.ML
    ID:         $Id$
    Author:     Tobias Nipkow
    Copyright   1997 TUM

Proves that 'a discr is a po
*)

goalw thy [less_discr_def] "(x::('a::term)discr) << x";
by (rtac refl 1);
qed "less_discr_refl";

goalw thy [less_discr_def]
  "!!x. [| (x::('a::term)discr) << y; y << z |] ==> x <<  z";
by (etac trans 1);
by (assume_tac 1);
qed "less_discr_trans";

goalw thy [less_discr_def]
  "!!x. [| (x::('a::term)discr) << y; y << x |] ==> x=y";
by (assume_tac 1);
qed "less_discr_antisym";