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
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/trancl.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Transitive closure of a relation
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
rtrancl is refl/transitive closure;  trancl is transitive closure
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
Trancl = Lfp + Prod + 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
    trans   :: "('a * 'a)set => bool" 	(*transitivity predicate*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
    id	    :: "('a * 'a)set"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
    rtrancl :: "('a * 'a)set => ('a * 'a)set"	("(_^*)" [100] 100)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
    trancl  :: "('a * 'a)set => ('a * 'a)set"	("(_^+)" [100] 100)  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
    O	    :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
defs   
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    19
trans_def	"trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
comp_def	(*composition of relations*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    21
		"r O s == {xz. ? x y z. xz = (x,z) & (x,y):s & (y,z):r}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
id_def		(*the identity relation*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    23
		"id == {p. ? x. p = (x,x)}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
rtrancl_def	"r^* == lfp(%s. id Un (r O s))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
trancl_def	"r^+ == r O rtrancl(r)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
end