# HG changeset patch # User paulson # Date 1046252656 -3600 # Node ID 7f8c1b533e8b254a35947006ea945bbe564f0398 # Parent af0218406395b442789f636fbf31ec08885304f8 some x-symbols and some new lemmas diff -r af0218406395 -r 7f8c1b533e8b src/HOL/Relation.thy --- a/src/HOL/Relation.thy Tue Feb 25 19:08:15 2003 +0100 +++ b/src/HOL/Relation.thy Wed Feb 26 10:44:16 2003 +0100 @@ -27,7 +27,7 @@ "Id == {p. EX x. p = (x,x)}" diag :: "'a set => ('a * 'a) set" -- {* diagonal: identity over a set *} - "diag A == UN x:A. {(x,x)}" + "diag A == \x\A. {(x,x)}" Domain :: "('a * 'b) set => 'a set" "Domain r == {x. EX y. (x,y):r}" @@ -36,7 +36,7 @@ "Range r == Domain(r^-1)" Field :: "('a * 'a) set => 'a set" - "Field r == Domain r Un Range r" + "Field r == Domain r \ Range r" refl :: "['a set, ('a * 'a) set] => bool" -- {* reflexivity over a set *} "refl A r == r \ A \ A & (ALL x: A. (x,x) : r)" @@ -232,16 +232,16 @@ lemma Domain_diag [simp]: "Domain (diag A) = A" by blast -lemma Domain_Un_eq: "Domain(A Un B) = Domain(A) Un Domain(B)" +lemma Domain_Un_eq: "Domain(A \ B) = Domain(A) \ Domain(B)" by blast -lemma Domain_Int_subset: "Domain(A Int B) \ Domain(A) Int Domain(B)" +lemma Domain_Int_subset: "Domain(A \ B) \ Domain(A) \ Domain(B)" by blast lemma Domain_Diff_subset: "Domain(A) - Domain(B) \ Domain(A - B)" by blast -lemma Domain_Union: "Domain (Union S) = (UN A:S. Domain A)" +lemma Domain_Union: "Domain (Union S) = (\A\S. Domain A)" by blast lemma Domain_mono: "r \ s ==> Domain r \ Domain s" @@ -271,16 +271,16 @@ lemma Range_diag [simp]: "Range (diag A) = A" by auto -lemma Range_Un_eq: "Range(A Un B) = Range(A) Un Range(B)" +lemma Range_Un_eq: "Range(A \ B) = Range(A) \ Range(B)" by blast -lemma Range_Int_subset: "Range(A Int B) \ Range(A) Int Range(B)" +lemma Range_Int_subset: "Range(A \ B) \ Range(A) \ Range(B)" by blast lemma Range_Diff_subset: "Range(A) - Range(B) \ Range(A - B)" by blast -lemma Range_Union: "Range (Union S) = (UN A:S. Range A)" +lemma Range_Union: "Range (Union S) = (\A\S. Range A)" by blast @@ -312,13 +312,17 @@ lemma Image_Id [simp]: "Id `` A = A" by blast -lemma Image_diag [simp]: "diag A `` B = A Int B" +lemma Image_diag [simp]: "diag A `` B = A \ B" + by blast + +lemma Image_Int_subset: "R `` (A \ B) \ R `` A \ R `` B" by blast -lemma Image_Int_subset: "R `` (A Int B) \ R `` A Int R `` B" - by blast +lemma Image_Int_eq: + "single_valued (converse R) ==> R `` (A \ B) = R `` A \ R `` B" + by (simp add: single_valued_def, blast) -lemma Image_Un: "R `` (A Un B) = R `` A Un R `` B" +lemma Image_Un: "R `` (A \ B) = R `` A \ R `` B" by blast lemma Un_Image: "(R \ S) `` A = R `` A \ S `` A" @@ -327,19 +331,26 @@ lemma Image_subset: "r \ A \ B ==> r``C \ B" by (rules intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2) -lemma Image_eq_UN: "r``B = (UN y: B. r``{y})" +lemma Image_eq_UN: "r``B = (\y\ B. r``{y})" -- {* NOT suitable for rewriting *} by blast lemma Image_mono: "r' \ r ==> A' \ A ==> (r' `` A') \ (r `` A)" by blast -lemma Image_UN: "(r `` (UNION A B)) = (UN x:A.(r `` (B x)))" +lemma Image_UN: "(r `` (UNION A B)) = (\x\A. r `` (B x))" + by blast + +lemma Image_INT_subset: "(r `` INTER A B) \ (\x\A. r `` (B x))" by blast -lemma Image_INT_subset: "(r `` (INTER A B)) \ (INT x:A.(r `` (B x)))" - -- {* Converse inclusion fails *} - by blast +text{*Converse inclusion requires some assumptions*} +lemma Image_INT_eq: + "[|single_valued (r\); A\{}|] ==> r `` INTER A B = (\x\A. r `` B x)" +apply (rule equalityI) + apply (rule Image_INT_subset) +apply (simp add: single_valued_def, blast) +done lemma Image_subset_eq: "(r``A \ B) = (A \ - ((r^-1) `` (-B)))" by blast