src/HOLCF/Cont.thy
author regensbu
Fri, 06 Oct 1995 17:25:24 +0100
changeset 1274 ea0668a1c0ba
parent 1168 74be52691d62
child 1479 21eb5e156d91
permissions -rw-r--r--
added 8bit pragmas added directory ax_ops for sections axioms and ops added directory domain for sections domain and generated this is the type definition package of David Oheimb

(*  Title: 	HOLCF/cont.thy
    ID:         $Id$
    Author: 	Franz Regensburger
    Copyright   1993 Technische Universitaet Muenchen

    Results about continuity and monotonicity
*)

Cont = Fun3 +

(* 

   Now we change the default class! Form now on all untyped typevariables are
   of default class pcpo

*)


default pcpo

consts  
	monofun :: "('a::po => 'b::po) => bool"	(* monotonicity    *)
	contlub	:: "('a => 'b) => bool"		(* first cont. def *)
	cont	:: "('a => 'b) => bool"		(* secnd cont. def *)

defs 

monofun		"monofun(f) == ! x y. x << y --> f(x) << f(y)"

contlub		"contlub(f) == ! Y. is_chain(Y) --> 
				f(lub(range(Y))) = lub(range(% i.f(Y(i))))"

cont		"cont(f)   == ! Y. is_chain(Y) --> 
				range(% i.f(Y(i))) <<| f(lub(range(Y)))"

(* ------------------------------------------------------------------------ *)
(* the main purpose of cont.thy is to show:                                 *)
(*              monofun(f) & contlub(f)  <==> cont(f)                       *)
(* ------------------------------------------------------------------------ *)

end