src/HOL/Relation.thy
changeset 7912 0e42be14f136
parent 7014 11ee650edcd2
child 8268 722074b93cdd
equal deleted inserted replaced
7911:b8dee46d778a 7912:0e42be14f136
     4     Copyright   1996  University of Cambridge
     4     Copyright   1996  University of Cambridge
     5 *)
     5 *)
     6 
     6 
     7 Relation = Prod +
     7 Relation = Prod +
     8 
     8 
     9 consts
       
    10   O            :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
       
    11   converse     :: "('a*'b) set => ('b*'a) set"     ("(_^-1)" [1000] 999)
       
    12   "^^"         :: "[('a*'b) set,'a set] => 'b set" (infixl 90)
       
    13   
       
    14 defs
       
    15   comp_def         "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
       
    16   converse_def     "r^-1 == {(y,x). (x,y):r}"
       
    17   Image_def        "r ^^ s == {y. ? x:s. (x,y):r}"
       
    18   
       
    19 constdefs
     9 constdefs
    20   Id     :: "('a * 'a)set"                 (*the identity relation*)
    10   converse :: "('a*'b) set => ('b*'a) set"               ("(_^-1)" [1000] 999)
    21       "Id == {p. ? x. p = (x,x)}"
    11     "r^-1 == {(y,x). (x,y):r}"
    22 
    12 
    23   diag   :: "'a set => ('a * 'a)set"
    13   comp  :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set"  (infixr "O" 60)
       
    14     "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
       
    15 
       
    16   Image :: "[('a*'b) set,'a set] => 'b set"                (infixl "^^" 90)
       
    17     "r ^^ s == {y. ? x:s. (x,y):r}"
       
    18 
       
    19   Id    :: "('a * 'a)set"                            (*the identity relation*)
       
    20     "Id == {p. ? x. p = (x,x)}"
       
    21 
       
    22   diag  :: "'a set => ('a * 'a)set"          (*diagonal: identity over a set*)
    24     "diag(A) == UN x:A. {(x,x)}"
    23     "diag(A) == UN x:A. {(x,x)}"
    25   
    24   
    26   Domain :: "('a*'b) set => 'a set"
    25   Domain :: "('a*'b) set => 'a set"
    27     "Domain(r) == {x. ? y. (x,y):r}"
    26     "Domain(r) == {x. ? y. (x,y):r}"
    28 
    27