wenzelm@10358: (* Title: HOL/Relation.thy nipkow@1128: ID: $Id$ paulson@1983: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@1983: Copyright 1996 University of Cambridge nipkow@1128: *) nipkow@1128: berghofe@12905: header {* Relations *} berghofe@12905: nipkow@15131: theory Relation nipkow@15131: import Product_Type nipkow@15131: begin paulson@5978: wenzelm@12913: subsection {* Definitions *} wenzelm@12913: paulson@5978: constdefs wenzelm@10358: converse :: "('a * 'b) set => ('b * 'a) set" ("(_^-1)" [1000] 999) wenzelm@10358: "r^-1 == {(y, x). (x, y) : r}" wenzelm@10358: syntax (xsymbols) berghofe@12905: converse :: "('a * 'b) set => ('b * 'a) set" ("(_\)" [1000] 999) paulson@7912: wenzelm@10358: constdefs nipkow@12487: rel_comp :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set" (infixr "O" 60) wenzelm@12913: "r O s == {(x,z). EX y. (x, y) : s & (y, z) : r}" wenzelm@12913: oheimb@11136: Image :: "[('a * 'b) set, 'a set] => 'b set" (infixl "``" 90) wenzelm@12913: "r `` s == {y. EX x:s. (x,y):r}" paulson@7912: berghofe@12905: Id :: "('a * 'a) set" -- {* the identity relation *} wenzelm@12913: "Id == {p. EX x. p = (x,x)}" paulson@7912: berghofe@12905: diag :: "'a set => ('a * 'a) set" -- {* diagonal: identity over a set *} paulson@13830: "diag A == \x\A. {(x,x)}" wenzelm@12913: oheimb@11136: Domain :: "('a * 'b) set => 'a set" wenzelm@12913: "Domain r == {x. EX y. (x,y):r}" paulson@5978: oheimb@11136: Range :: "('a * 'b) set => 'b set" wenzelm@12913: "Range r == Domain(r^-1)" paulson@5978: oheimb@11136: Field :: "('a * 'a) set => 'a set" paulson@13830: "Field r == Domain r \ Range r" paulson@10786: berghofe@12905: refl :: "['a set, ('a * 'a) set] => bool" -- {* reflexivity over a set *} wenzelm@12913: "refl A r == r \ A \ A & (ALL x: A. (x,x) : r)" paulson@6806: berghofe@12905: sym :: "('a * 'a) set => bool" -- {* symmetry predicate *} wenzelm@12913: "sym r == ALL x y. (x,y): r --> (y,x): r" paulson@6806: berghofe@12905: antisym:: "('a * 'a) set => bool" -- {* antisymmetry predicate *} wenzelm@12913: "antisym r == ALL x y. (x,y):r --> (y,x):r --> x=y" paulson@6806: berghofe@12905: trans :: "('a * 'a) set => bool" -- {* transitivity predicate *} wenzelm@12913: "trans r == (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)" paulson@5978: oheimb@11136: single_valued :: "('a * 'b) set => bool" wenzelm@12913: "single_valued r == ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z)" berghofe@7014: oheimb@11136: inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" wenzelm@12913: "inv_image r f == {(x, y). (f x, f y) : r}" oheimb@11136: paulson@6806: syntax berghofe@12905: reflexive :: "('a * 'a) set => bool" -- {* reflexivity over a type *} paulson@6806: translations paulson@6806: "reflexive" == "refl UNIV" paulson@6806: berghofe@12905: wenzelm@12913: subsection {* The identity relation *} berghofe@12905: berghofe@12905: lemma IdI [intro]: "(a, a) : Id" berghofe@12905: by (simp add: Id_def) berghofe@12905: berghofe@12905: lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P" berghofe@12905: by (unfold Id_def) (rules elim: CollectE) berghofe@12905: berghofe@12905: lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)" berghofe@12905: by (unfold Id_def) blast berghofe@12905: berghofe@12905: lemma reflexive_Id: "reflexive Id" berghofe@12905: by (simp add: refl_def) berghofe@12905: berghofe@12905: lemma antisym_Id: "antisym Id" berghofe@12905: -- {* A strange result, since @{text Id} is also symmetric. *} berghofe@12905: by (simp add: antisym_def) berghofe@12905: berghofe@12905: lemma trans_Id: "trans Id" berghofe@12905: by (simp add: trans_def) berghofe@12905: berghofe@12905: wenzelm@12913: subsection {* Diagonal: identity over a set *} berghofe@12905: paulson@13812: lemma diag_empty [simp]: "diag {} = {}" paulson@13812: by (simp add: diag_def) paulson@13812: berghofe@12905: lemma diag_eqI: "a = b ==> a : A ==> (a, b) : diag A" berghofe@12905: by (simp add: diag_def) berghofe@12905: berghofe@12905: lemma diagI [intro!]: "a : A ==> (a, a) : diag A" berghofe@12905: by (rule diag_eqI) (rule refl) berghofe@12905: berghofe@12905: lemma diagE [elim!]: berghofe@12905: "c : diag A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P" wenzelm@12913: -- {* The general elimination rule. *} berghofe@12905: by (unfold diag_def) (rules elim!: UN_E singletonE) berghofe@12905: berghofe@12905: lemma diag_iff: "((x, y) : diag A) = (x = y & x : A)" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma diag_subset_Times: "diag A \ A \ A" berghofe@12905: by blast berghofe@12905: berghofe@12905: berghofe@12905: subsection {* Composition of two relations *} berghofe@12905: wenzelm@12913: lemma rel_compI [intro]: berghofe@12905: "(a, b) : s ==> (b, c) : r ==> (a, c) : r O s" berghofe@12905: by (unfold rel_comp_def) blast berghofe@12905: wenzelm@12913: lemma rel_compE [elim!]: "xz : r O s ==> berghofe@12905: (!!x y z. xz = (x, z) ==> (x, y) : s ==> (y, z) : r ==> P) ==> P" berghofe@12905: by (unfold rel_comp_def) (rules elim!: CollectE splitE exE conjE) berghofe@12905: berghofe@12905: lemma rel_compEpair: berghofe@12905: "(a, c) : r O s ==> (!!y. (a, y) : s ==> (y, c) : r ==> P) ==> P" berghofe@12905: by (rules elim: rel_compE Pair_inject ssubst) berghofe@12905: berghofe@12905: lemma R_O_Id [simp]: "R O Id = R" berghofe@12905: by fast berghofe@12905: berghofe@12905: lemma Id_O_R [simp]: "Id O R = R" berghofe@12905: by fast berghofe@12905: berghofe@12905: lemma O_assoc: "(R O S) O T = R O (S O T)" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma trans_O_subset: "trans r ==> r O r \ r" berghofe@12905: by (unfold trans_def) blast berghofe@12905: wenzelm@12913: lemma rel_comp_mono: "r' \ r ==> s' \ s ==> (r' O s') \ (r O s)" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma rel_comp_subset_Sigma: wenzelm@12913: "s \ A \ B ==> r \ B \ C ==> (r O s) \ A \ C" berghofe@12905: by blast berghofe@12905: wenzelm@12913: wenzelm@12913: subsection {* Reflexivity *} wenzelm@12913: wenzelm@12913: lemma reflI: "r \ A \ A ==> (!!x. x : A ==> (x, x) : r) ==> refl A r" berghofe@12905: by (unfold refl_def) (rules intro!: ballI) berghofe@12905: berghofe@12905: lemma reflD: "refl A r ==> a : A ==> (a, a) : r" berghofe@12905: by (unfold refl_def) blast berghofe@12905: wenzelm@12913: wenzelm@12913: subsection {* Antisymmetry *} berghofe@12905: berghofe@12905: lemma antisymI: berghofe@12905: "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r" berghofe@12905: by (unfold antisym_def) rules berghofe@12905: berghofe@12905: lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b" berghofe@12905: by (unfold antisym_def) rules berghofe@12905: wenzelm@12913: wenzelm@12913: subsection {* Transitivity *} berghofe@12905: berghofe@12905: lemma transI: berghofe@12905: "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r" berghofe@12905: by (unfold trans_def) rules berghofe@12905: berghofe@12905: lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r" berghofe@12905: by (unfold trans_def) rules berghofe@12905: berghofe@12905: wenzelm@12913: subsection {* Converse *} wenzelm@12913: wenzelm@12913: lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)" berghofe@12905: by (simp add: converse_def) berghofe@12905: nipkow@13343: lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1" berghofe@12905: by (simp add: converse_def) berghofe@12905: nipkow@13343: lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r" berghofe@12905: by (simp add: converse_def) berghofe@12905: berghofe@12905: lemma converseE [elim!]: berghofe@12905: "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P" wenzelm@12913: -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *} berghofe@12905: by (unfold converse_def) (rules elim!: CollectE splitE bexE) berghofe@12905: berghofe@12905: lemma converse_converse [simp]: "(r^-1)^-1 = r" berghofe@12905: by (unfold converse_def) blast berghofe@12905: berghofe@12905: lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma converse_Id [simp]: "Id^-1 = Id" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma converse_diag [simp]: "(diag A)^-1 = diag A" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma refl_converse: "refl A r ==> refl A (converse r)" berghofe@12905: by (unfold refl_def) blast berghofe@12905: berghofe@12905: lemma antisym_converse: "antisym (converse r) = antisym r" berghofe@12905: by (unfold antisym_def) blast berghofe@12905: berghofe@12905: lemma trans_converse: "trans (converse r) = trans r" berghofe@12905: by (unfold trans_def) blast berghofe@12905: wenzelm@12913: berghofe@12905: subsection {* Domain *} berghofe@12905: berghofe@12905: lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)" berghofe@12905: by (unfold Domain_def) blast berghofe@12905: berghofe@12905: lemma DomainI [intro]: "(a, b) : r ==> a : Domain r" berghofe@12905: by (rules intro!: iffD2 [OF Domain_iff]) berghofe@12905: berghofe@12905: lemma DomainE [elim!]: berghofe@12905: "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P" berghofe@12905: by (rules dest!: iffD1 [OF Domain_iff]) berghofe@12905: berghofe@12905: lemma Domain_empty [simp]: "Domain {} = {}" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Domain_Id [simp]: "Domain Id = UNIV" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Domain_diag [simp]: "Domain (diag A) = A" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Domain_Un_eq: "Domain(A \ B) = Domain(A) \ Domain(B)" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Domain_Int_subset: "Domain(A \ B) \ Domain(A) \ Domain(B)" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma Domain_Diff_subset: "Domain(A) - Domain(B) \ Domain(A - B)" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Domain_Union: "Domain (Union S) = (\A\S. Domain A)" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma Domain_mono: "r \ s ==> Domain r \ Domain s" berghofe@12905: by blast berghofe@12905: berghofe@12905: berghofe@12905: subsection {* Range *} berghofe@12905: berghofe@12905: lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)" berghofe@12905: by (simp add: Domain_def Range_def) berghofe@12905: berghofe@12905: lemma RangeI [intro]: "(a, b) : r ==> b : Range r" berghofe@12905: by (unfold Range_def) (rules intro!: converseI DomainI) berghofe@12905: berghofe@12905: lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P" berghofe@12905: by (unfold Range_def) (rules elim!: DomainE dest!: converseD) berghofe@12905: berghofe@12905: lemma Range_empty [simp]: "Range {} = {}" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Range_Id [simp]: "Range Id = UNIV" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Range_diag [simp]: "Range (diag A) = A" berghofe@12905: by auto berghofe@12905: paulson@13830: lemma Range_Un_eq: "Range(A \ B) = Range(A) \ Range(B)" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Range_Int_subset: "Range(A \ B) \ Range(A) \ Range(B)" berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma Range_Diff_subset: "Range(A) - Range(B) \ Range(A - B)" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Range_Union: "Range (Union S) = (\A\S. Range A)" berghofe@12905: by blast berghofe@12905: berghofe@12905: berghofe@12905: subsection {* Image of a set under a relation *} berghofe@12905: wenzelm@12913: lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)" berghofe@12905: by (simp add: Image_def) berghofe@12905: wenzelm@12913: lemma Image_singleton: "r``{a} = {b. (a, b) : r}" berghofe@12905: by (simp add: Image_def) berghofe@12905: wenzelm@12913: lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)" berghofe@12905: by (rule Image_iff [THEN trans]) simp berghofe@12905: wenzelm@12913: lemma ImageI [intro]: "(a, b) : r ==> a : A ==> b : r``A" berghofe@12905: by (unfold Image_def) blast berghofe@12905: berghofe@12905: lemma ImageE [elim!]: wenzelm@12913: "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P" berghofe@12905: by (unfold Image_def) (rules elim!: CollectE bexE) berghofe@12905: berghofe@12905: lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A" berghofe@12905: -- {* This version's more effective when we already have the required @{text a} *} berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Image_empty [simp]: "R``{} = {}" berghofe@12905: by blast berghofe@12905: berghofe@12905: lemma Image_Id [simp]: "Id `` A = A" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Image_diag [simp]: "diag A `` B = A \ B" paulson@13830: by blast paulson@13830: paulson@13830: lemma Image_Int_subset: "R `` (A \ B) \ R `` A \ R `` B" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Image_Int_eq: paulson@13830: "single_valued (converse R) ==> R `` (A \ B) = R `` A \ R `` B" paulson@13830: by (simp add: single_valued_def, blast) berghofe@12905: paulson@13830: lemma Image_Un: "R `` (A \ B) = R `` A \ R `` B" berghofe@12905: by blast berghofe@12905: paulson@13812: lemma Un_Image: "(R \ S) `` A = R `` A \ S `` A" paulson@13812: by blast paulson@13812: wenzelm@12913: lemma Image_subset: "r \ A \ B ==> r``C \ B" berghofe@12905: by (rules intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2) berghofe@12905: paulson@13830: lemma Image_eq_UN: "r``B = (\y\ B. r``{y})" berghofe@12905: -- {* NOT suitable for rewriting *} berghofe@12905: by blast berghofe@12905: wenzelm@12913: lemma Image_mono: "r' \ r ==> A' \ A ==> (r' `` A') \ (r `` A)" berghofe@12905: by blast berghofe@12905: paulson@13830: lemma Image_UN: "(r `` (UNION A B)) = (\x\A. r `` (B x))" paulson@13830: by blast paulson@13830: paulson@13830: lemma Image_INT_subset: "(r `` INTER A B) \ (\x\A. r `` (B x))" berghofe@12905: by blast berghofe@12905: paulson@13830: text{*Converse inclusion requires some assumptions*} paulson@13830: lemma Image_INT_eq: paulson@13830: "[|single_valued (r\); A\{}|] ==> r `` INTER A B = (\x\A. r `` B x)" paulson@13830: apply (rule equalityI) paulson@13830: apply (rule Image_INT_subset) paulson@13830: apply (simp add: single_valued_def, blast) paulson@13830: done berghofe@12905: wenzelm@12913: lemma Image_subset_eq: "(r``A \ B) = (A \ - ((r^-1) `` (-B)))" berghofe@12905: by blast berghofe@12905: berghofe@12905: wenzelm@12913: subsection {* Single valued relations *} wenzelm@12913: wenzelm@12913: lemma single_valuedI: berghofe@12905: "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r" berghofe@12905: by (unfold single_valued_def) berghofe@12905: berghofe@12905: lemma single_valuedD: berghofe@12905: "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z" berghofe@12905: by (simp add: single_valued_def) berghofe@12905: berghofe@12905: berghofe@12905: subsection {* Graphs given by @{text Collect} *} berghofe@12905: berghofe@12905: lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}" berghofe@12905: by auto berghofe@12905: berghofe@12905: lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}" berghofe@12905: by auto berghofe@12905: berghofe@12905: lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}" berghofe@12905: by auto berghofe@12905: berghofe@12905: wenzelm@12913: subsection {* Inverse image *} berghofe@12905: wenzelm@12913: lemma trans_inv_image: "trans r ==> trans (inv_image r f)" berghofe@12905: apply (unfold trans_def inv_image_def) berghofe@12905: apply (simp (no_asm)) berghofe@12905: apply blast berghofe@12905: done berghofe@12905: nipkow@1128: end