clasohm@1475: (* Title: Relation.thy nipkow@1128: ID: $Id$ clasohm@1475: Author: Riccardo Mattolini, Dip. Sistemi e Informatica clasohm@1475: and Lawrence C Paulson, Cambridge University Computer Laboratory nipkow@1128: Copyright 1994 Universita' di Firenze nipkow@1128: Copyright 1993 University of Cambridge nipkow@1128: *) nipkow@1128: nipkow@1128: Relation = Prod + nipkow@1128: consts clasohm@1475: id :: "('a * 'a)set" (*the identity relation*) clasohm@1475: O :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60) clasohm@1475: trans :: "('a * 'a)set => bool" (*transitivity predicate*) nipkow@1454: converse :: "('a * 'b)set => ('b * 'a)set" nipkow@1454: "^^" :: "[('a * 'b) set, 'a set] => 'b set" (infixl 90) nipkow@1454: Domain :: "('a * 'b) set => 'a set" nipkow@1454: Range :: "('a * 'b) set => 'b set" nipkow@1128: defs clasohm@1475: id_def "id == {p. ? x. p = (x,x)}" clasohm@1475: comp_def "r O s == {(x,z). ? y. (x,y):s & (y,z):r}" clasohm@1475: trans_def "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)" nipkow@1454: converse_def "converse(r) == {(y,x). (x,y):r}" nipkow@1454: Domain_def "Domain(r) == {x. ? y. (x,y):r}" nipkow@1128: Range_def "Range(r) == Domain(converse(r))" nipkow@1128: Image_def "r ^^ s == {y. y:Range(r) & (? x:s. (x,y):r)}" nipkow@1128: end