diff -r 8fe3e66abf0c -r c22b85994e17 src/HOLCF/porder.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/porder.thy Wed Jan 19 17:35:01 1994 +0100 @@ -0,0 +1,71 @@ +(* Title: HOLCF/porder.thy + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Definition of class porder (partial order) + +The prototype theory for this class is void.thy + +*) + +Porder = Void + + +(* Introduction of new class. The witness is type void. *) + +classes po < term + +(* default type is still term ! *) +(* void is the prototype in po *) + +arities void :: po + +consts "<<" :: "['a,'a::po] => bool" (infixl 55) + + "<|" :: "['a set,'a::po] => bool" (infixl 55) + "<<|" :: "['a set,'a::po] => bool" (infixl 55) + lub :: "'a set => 'a::po" + is_tord :: "'a::po set => bool" + is_chain :: "(nat=>'a::po) => bool" + max_in_chain :: "[nat,nat=>'a::po]=>bool" + finite_chain :: "(nat=>'a::po)=>bool" + +rules + +(* class axioms: justification is theory Void *) + +refl_less "x << x" + (* witness refl_less_void *) + +antisym_less "[|x< x = y" + (* witness antisym_less_void *) + +trans_less "[|x< x<bool = less_void" + +(* class definitions *) + +is_ub "S <| x == ! y.y:S --> y< x << u)" + +lub "lub(S) = (@x. S <<| x)" + +(* Arbitrary chains are total orders *) +is_tord "is_tord(S) == ! x y. x:S & y:S --> (x< C(i) = C(j)" + +finite_chain_def "finite_chain(C) == is_chain(C) & (? i. max_in_chain(i,C))" + +end