src/CCL/Trancl.thy
author kleing
Mon, 12 May 2003 14:49:08 +0200
changeset 14012 9d1f027eb4e8
parent 1474 3f7d67927fe2
child 17456 bcf7544875b2
permissions -rw-r--r--
be nice on sunbroy2

(*  Title:      CCL/trancl.thy
    ID:         $Id$
    Author:     Martin Coen, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

Transitive closure of a relation
*)

Trancl = CCL +

consts
    trans   :: "i set => o"                   (*transitivity predicate*)
    id      :: "i set"
    rtrancl :: "i set => i set"               ("(_^*)" [100] 100)
    trancl  :: "i set => i set"               ("(_^+)" [100] 100)  
    O       :: "[i set,i set] => i set"       (infixr 60)

rules   

trans_def       "trans(r) == (ALL x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)"
comp_def        (*composition of relations*)
                "r O s == {xz. EX x y z. xz = <x,z> & <x,y>:s & <y,z>:r}"
id_def          (*the identity relation*)
                "id == {p. EX x. p = <x,x>}"
rtrancl_def     "r^* == lfp(%s. id Un (r O s))"
trancl_def      "r^+ == r O rtrancl(r)"

end