src/HOL/Relation.thy
changeset 1454 d0266c81a85e
parent 1128 64b30e3cc6d4
child 1475 7f5a4cd08209
--- a/src/HOL/Relation.thy	Fri Jan 26 13:43:36 1996 +0100
+++ b/src/HOL/Relation.thy	Fri Jan 26 20:25:39 1996 +0100
@@ -11,17 +11,16 @@
     id	        :: "('a * 'a)set"               (*the identity relation*)
     O	        :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
     trans       :: "('a * 'a)set => bool" 	(*transitivity predicate*)
-    converse    :: "('a*'a) set => ('a*'a) set"
-    "^^"        :: "[('a*'a) set,'a set] => 'a set" (infixl 90)
-    Domain      :: "('a*'a) set => 'a set"
-    Range       :: "('a*'a) set => 'a set"
+    converse    :: "('a * 'b)set => ('b * 'a)set"
+    "^^"        :: "[('a * 'b) set, 'a set] => 'b set" (infixl 90)
+    Domain      :: "('a * 'b) set => 'a set"
+    Range       :: "('a * 'b) set => 'b set"
 defs
     id_def	"id == {p. ? x. p = (x,x)}"
-    comp_def	(*composition of relations*)
-		"r O s == {xz. ? x y z. xz = (x,z) & (x,y):s & (y,z):r}"
+    comp_def	"r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
     trans_def	  "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
-    converse_def  "converse(r) == {z. (? w:r. ? x y. w=(x,y) & z=(y,x))}"
-    Domain_def    "Domain(r) == {z. ! x. (z=x --> (? y. (x,y):r))}"
+    converse_def  "converse(r) == {(y,x). (x,y):r}"
+    Domain_def    "Domain(r) == {x. ? y. (x,y):r}"
     Range_def     "Range(r) == Domain(converse(r))"
     Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. (x,y):r)}"
 end