src/HOL/Relation.thy
changeset 1475 7f5a4cd08209
parent 1454 d0266c81a85e
child 1695 0f9b9eda2a2c
equal deleted inserted replaced
1474:3f7d67927fe2 1475:7f5a4cd08209
     1 (*  Title: 	Relation.thy
     1 (*  Title:      Relation.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author: 	Riccardo Mattolini, Dip. Sistemi e Informatica
     3     Author:     Riccardo Mattolini, Dip. Sistemi e Informatica
     4         and	Lawrence C Paulson, Cambridge University Computer Laboratory
     4         and     Lawrence C Paulson, Cambridge University Computer Laboratory
     5     Copyright   1994 Universita' di Firenze
     5     Copyright   1994 Universita' di Firenze
     6     Copyright   1993  University of Cambridge
     6     Copyright   1993  University of Cambridge
     7 *)
     7 *)
     8 
     8 
     9 Relation = Prod +
     9 Relation = Prod +
    10 consts
    10 consts
    11     id	        :: "('a * 'a)set"               (*the identity relation*)
    11     id          :: "('a * 'a)set"               (*the identity relation*)
    12     O	        :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
    12     O           :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
    13     trans       :: "('a * 'a)set => bool" 	(*transitivity predicate*)
    13     trans       :: "('a * 'a)set => bool"       (*transitivity predicate*)
    14     converse    :: "('a * 'b)set => ('b * 'a)set"
    14     converse    :: "('a * 'b)set => ('b * 'a)set"
    15     "^^"        :: "[('a * 'b) set, 'a set] => 'b set" (infixl 90)
    15     "^^"        :: "[('a * 'b) set, 'a set] => 'b set" (infixl 90)
    16     Domain      :: "('a * 'b) set => 'a set"
    16     Domain      :: "('a * 'b) set => 'a set"
    17     Range       :: "('a * 'b) set => 'b set"
    17     Range       :: "('a * 'b) set => 'b set"
    18 defs
    18 defs
    19     id_def	"id == {p. ? x. p = (x,x)}"
    19     id_def      "id == {p. ? x. p = (x,x)}"
    20     comp_def	"r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
    20     comp_def    "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
    21     trans_def	  "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
    21     trans_def     "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
    22     converse_def  "converse(r) == {(y,x). (x,y):r}"
    22     converse_def  "converse(r) == {(y,x). (x,y):r}"
    23     Domain_def    "Domain(r) == {x. ? y. (x,y):r}"
    23     Domain_def    "Domain(r) == {x. ? y. (x,y):r}"
    24     Range_def     "Range(r) == Domain(converse(r))"
    24     Range_def     "Range(r) == Domain(converse(r))"
    25     Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. (x,y):r)}"
    25     Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. (x,y):r)}"
    26 end
    26 end