src/HOLCF/porder.thy
changeset 243 c22b85994e17
child 297 5ef75ff3baeb
equal deleted inserted replaced
242:8fe3e66abf0c 243:c22b85994e17
       
     1 (*  Title: 	HOLCF/porder.thy
       
     2     ID:         $Id$
       
     3     Author: 	Franz Regensburger
       
     4     Copyright   1993 Technische Universitaet Muenchen
       
     5 
       
     6 Definition of class porder (partial order)
       
     7 
       
     8 The prototype theory for this class is void.thy 
       
     9 
       
    10 *)
       
    11 
       
    12 Porder = Void +
       
    13 
       
    14 (* Introduction of new class. The witness is type void. *)
       
    15 
       
    16 classes po < term
       
    17 
       
    18 (* default type is still term ! *)
       
    19 (* void is the prototype in po *)
       
    20 
       
    21 arities void :: po
       
    22 
       
    23 consts	"<<"	::	"['a,'a::po] => bool"	(infixl 55)
       
    24 
       
    25 	"<|"	::	"['a set,'a::po] => bool"	(infixl 55)
       
    26 	"<<|"	::	"['a set,'a::po] => bool"	(infixl 55)
       
    27 	lub	::	"'a set => 'a::po"
       
    28 	is_tord	::	"'a::po set => bool"
       
    29 	is_chain ::	"(nat=>'a::po) => bool"
       
    30 	max_in_chain :: "[nat,nat=>'a::po]=>bool"
       
    31 	finite_chain :: "(nat=>'a::po)=>bool"
       
    32 
       
    33 rules
       
    34 
       
    35 (* class axioms: justification is theory Void *)
       
    36 
       
    37 refl_less	"x << x"	
       
    38 				(* witness refl_less_void    *)
       
    39 
       
    40 antisym_less	"[|x<<y ; y<<x |] ==> x = y"	
       
    41 				(* witness antisym_less_void *)
       
    42 
       
    43 trans_less	"[|x<<y ; y<<z |] ==> x<<z"
       
    44 				(* witness trans_less_void   *)
       
    45 
       
    46 (* instance of << for the prototype void *)
       
    47 
       
    48 inst_void_po	"(op <<)::[void,void]=>bool = less_void"
       
    49 
       
    50 (* class definitions *)
       
    51 
       
    52 is_ub		"S  <| x == ! y.y:S --> y<<x"
       
    53 is_lub		"S <<| x == S <| x & (! u. S <| u  --> x << u)"
       
    54 
       
    55 lub		"lub(S) = (@x. S <<| x)"
       
    56 
       
    57 (* Arbitrary chains are total orders    *)                  
       
    58 is_tord		"is_tord(S) == ! x y. x:S & y:S --> (x<<y | y<<x)"
       
    59 
       
    60 
       
    61 (* Here we use countable chains and I prefer to code them as functions! *)
       
    62 is_chain	"is_chain(F) == (! i.F(i) << F(Suc(i)))"
       
    63 
       
    64 
       
    65 (* finite chains, needed for monotony of continouous functions *)
       
    66 
       
    67 max_in_chain_def "max_in_chain(i,C) == ! j. i <= j --> C(i) = C(j)" 
       
    68 
       
    69 finite_chain_def "finite_chain(C) == is_chain(C) & (? i. max_in_chain(i,C))"
       
    70 
       
    71 end