src/HOL/Relation.thy
changeset 5978 fa2c2dd74f8c
parent 5608 a82a038a3e7a
child 6806 43c081a0858d
--- a/src/HOL/Relation.thy	Thu Nov 26 17:40:10 1998 +0100
+++ b/src/HOL/Relation.thy	Fri Nov 27 10:40:29 1998 +0100
@@ -5,22 +5,34 @@
 *)
 
 Relation = Prod +
+
 consts
-    Id          :: "('a * 'a)set"               (*the identity relation*)
-    O           :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
-    converse    :: "('a*'b) set => ('b*'a) set"     ("(_^-1)" [1000] 999)
-    "^^"        :: "[('a*'b) set,'a set] => 'b set" (infixl 90)
-    Domain      :: "('a*'b) set => 'a set"
-    Range       :: "('a*'b) set => 'b set"
-    trans       :: "('a * 'a)set => bool"       (*transitivity predicate*)
-    Univalent   :: "('a * 'b)set => bool"
+  O           :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
+  converse    :: "('a*'b) set => ('b*'a) set"     ("(_^-1)" [1000] 999)
+  "^^"        :: "[('a*'b) set,'a set] => 'b set" (infixl 90)
+  
 defs
-    Id_def        "Id == {p. ? x. p = (x,x)}"
-    comp_def      "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
-    converse_def   "r^-1 == {(y,x). (x,y):r}"
-    Domain_def    "Domain(r) == {x. ? y. (x,y):r}"
-    Range_def     "Range(r) == Domain(r^-1)"
-    Image_def     "r ^^ s == {y. ? x:s. (x,y):r}"
-    trans_def     "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
-    Univalent_def "Univalent r == !x y. (x,y):r --> (!z. (x,z):r --> y=z)"
+  comp_def      "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
+  converse_def  "r^-1 == {(y,x). (x,y):r}"
+  Image_def     "r ^^ s == {y. ? x:s. (x,y):r}"
+  
+constdefs
+  Id          :: "('a * 'a)set"               (*the identity relation*)
+      "Id == {p. ? x. p = (x,x)}"
+
+  diag   :: "'a set => ('a * 'a)set"
+    "diag(A) == UN x:A. {(x,x)}"
+  
+  Domain      :: "('a*'b) set => 'a set"
+    "Domain(r) == {x. ? y. (x,y):r}"
+
+  Range       :: "('a*'b) set => 'b set"
+    "Range(r) == Domain(r^-1)"
+
+  trans       :: "('a * 'a)set => bool"       (*transitivity predicate*)
+    "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
+
+  Univalent   :: "('a * 'b)set => bool"
+    "Univalent r == !x y. (x,y):r --> (!z. (x,z):r --> y=z)"
+
 end