src/HOL/Trancl.thy
author lcp
Tue, 25 Apr 1995 11:14:03 +0200
changeset 1072 0140ff702b23
parent 972 e61b058d58d2
child 1128 64b30e3cc6d4
permissions -rw-r--r--
updated version

(*  Title: 	HOL/trancl.thy
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge

Transitive closure of a relation

rtrancl is refl/transitive closure;  trancl is transitive closure
*)

Trancl = Lfp + Prod + 
consts
    trans   :: "('a * 'a)set => bool" 	(*transitivity predicate*)
    id	    :: "('a * 'a)set"
    rtrancl :: "('a * 'a)set => ('a * 'a)set"	("(_^*)" [100] 100)
    trancl  :: "('a * 'a)set => ('a * 'a)set"	("(_^+)" [100] 100)  
    O	    :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
defs   
trans_def	"trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
comp_def	(*composition of relations*)
		"r O s == {xz. ? x y z. xz = (x,z) & (x,y):s & (y,z):r}"
id_def		(*the identity relation*)
		"id == {p. ? x. p = (x,x)}"
rtrancl_def	"r^* == lfp(%s. id Un (r O s))"
trancl_def	"r^+ == r O rtrancl(r)"
end