src/HOL/Relation.thy
author bulwahn
Fri, 03 Dec 2010 08:40:47 +0100
changeset 41167 be80c93ac0a2
parent 36764 ef97c5006840
child 41304 dcec9bc71ee9
permissions -rw-r--r--
adding a nice definition of Id_on for quickcheck and nitpick
     1 (*  Title:      HOL/Relation.thy
     2     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     3     Copyright   1996  University of Cambridge
     4 *)
     5 
     6 header {* Relations *}
     7 
     8 theory Relation
     9 imports Datatype Finite_Set
    10 begin
    11 
    12 subsection {* Definitions *}
    13 
    14 definition
    15   converse :: "('a * 'b) set => ('b * 'a) set"
    16     ("(_^-1)" [1000] 999) where
    17   "r^-1 == {(y, x). (x, y) : r}"
    18 
    19 notation (xsymbols)
    20   converse  ("(_\<inverse>)" [1000] 999)
    21 
    22 definition
    23   rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
    24     (infixr "O" 75) where
    25   "r O s == {(x,z). EX y. (x, y) : r & (y, z) : s}"
    26 
    27 definition
    28   Image :: "[('a * 'b) set, 'a set] => 'b set"
    29     (infixl "``" 90) where
    30   "r `` s == {y. EX x:s. (x,y):r}"
    31 
    32 definition
    33   Id :: "('a * 'a) set" where -- {* the identity relation *}
    34   "Id == {p. EX x. p = (x,x)}"
    35 
    36 definition
    37   Id_on  :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
    38   "Id_on A == \<Union>x\<in>A. {(x,x)}"
    39 
    40 definition
    41   Domain :: "('a * 'b) set => 'a set" where
    42   "Domain r == {x. EX y. (x,y):r}"
    43 
    44 definition
    45   Range  :: "('a * 'b) set => 'b set" where
    46   "Range r == Domain(r^-1)"
    47 
    48 definition
    49   Field :: "('a * 'a) set => 'a set" where
    50   "Field r == Domain r \<union> Range r"
    51 
    52 definition
    53   refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
    54   "refl_on A r == r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
    55 
    56 abbreviation
    57   refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
    58   "refl == refl_on UNIV"
    59 
    60 definition
    61   sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
    62   "sym r == ALL x y. (x,y): r --> (y,x): r"
    63 
    64 definition
    65   antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
    66   "antisym r == ALL x y. (x,y):r --> (y,x):r --> x=y"
    67 
    68 definition
    69   trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
    70   "trans r == (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
    71 
    72 definition
    73 irrefl :: "('a * 'a) set => bool" where
    74 "irrefl r \<equiv> \<forall>x. (x,x) \<notin> r"
    75 
    76 definition
    77 total_on :: "'a set => ('a * 'a) set => bool" where
    78 "total_on A r \<equiv> \<forall>x\<in>A.\<forall>y\<in>A. x\<noteq>y \<longrightarrow> (x,y)\<in>r \<or> (y,x)\<in>r"
    79 
    80 abbreviation "total \<equiv> total_on UNIV"
    81 
    82 definition
    83   single_valued :: "('a * 'b) set => bool" where
    84   "single_valued r == ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z)"
    85 
    86 definition
    87   inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
    88   "inv_image r f == {(x, y). (f x, f y) : r}"
    89 
    90 
    91 subsection {* The identity relation *}
    92 
    93 lemma IdI [intro]: "(a, a) : Id"
    94 by (simp add: Id_def)
    95 
    96 lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
    97 by (unfold Id_def) (iprover elim: CollectE)
    98 
    99 lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)"
   100 by (unfold Id_def) blast
   101 
   102 lemma refl_Id: "refl Id"
   103 by (simp add: refl_on_def)
   104 
   105 lemma antisym_Id: "antisym Id"
   106   -- {* A strange result, since @{text Id} is also symmetric. *}
   107 by (simp add: antisym_def)
   108 
   109 lemma sym_Id: "sym Id"
   110 by (simp add: sym_def)
   111 
   112 lemma trans_Id: "trans Id"
   113 by (simp add: trans_def)
   114 
   115 
   116 subsection {* Diagonal: identity over a set *}
   117 
   118 lemma Id_on_empty [simp]: "Id_on {} = {}"
   119 by (simp add: Id_on_def) 
   120 
   121 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   122 by (simp add: Id_on_def)
   123 
   124 lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A"
   125 by (rule Id_on_eqI) (rule refl)
   126 
   127 lemma Id_onE [elim!]:
   128   "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
   129   -- {* The general elimination rule. *}
   130 by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
   131 
   132 lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
   133 by blast
   134 
   135 lemma Id_on_def'[nitpick_def, code]:
   136   "(Id_on (A :: 'a => bool)) = (%(x, y). x = y \<and> A x)"
   137 by (auto simp add: fun_eq_iff
   138   elim: Id_onE[unfolded mem_def] intro: Id_onI[unfolded mem_def])
   139 
   140 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   141 by blast
   142 
   143 
   144 subsection {* Composition of two relations *}
   145 
   146 lemma rel_compI [intro]:
   147   "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   148 by (unfold rel_comp_def) blast
   149 
   150 lemma rel_compE [elim!]: "xz : r O s ==>
   151   (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s  ==> P) ==> P"
   152 by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
   153 
   154 lemma rel_compEpair:
   155   "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P"
   156 by (iprover elim: rel_compE Pair_inject ssubst)
   157 
   158 lemma R_O_Id [simp]: "R O Id = R"
   159 by fast
   160 
   161 lemma Id_O_R [simp]: "Id O R = R"
   162 by fast
   163 
   164 lemma rel_comp_empty1[simp]: "{} O R = {}"
   165 by blast
   166 
   167 lemma rel_comp_empty2[simp]: "R O {} = {}"
   168 by blast
   169 
   170 lemma O_assoc: "(R O S) O T = R O (S O T)"
   171 by blast
   172 
   173 lemma trans_O_subset: "trans r ==> r O r \<subseteq> r"
   174 by (unfold trans_def) blast
   175 
   176 lemma rel_comp_mono: "r' \<subseteq> r ==> s' \<subseteq> s ==> (r' O s') \<subseteq> (r O s)"
   177 by blast
   178 
   179 lemma rel_comp_subset_Sigma:
   180     "r \<subseteq> A \<times> B ==> s \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
   181 by blast
   182 
   183 lemma rel_comp_distrib[simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)" 
   184 by auto
   185 
   186 lemma rel_comp_distrib2[simp]: "(S \<union> T) O R = (S O R) \<union> (T O R)"
   187 by auto
   188 
   189 lemma rel_comp_UNION_distrib: "s O UNION I r = UNION I (%i. s O r i)"
   190 by auto
   191 
   192 lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)"
   193 by auto
   194 
   195 
   196 subsection {* Reflexivity *}
   197 
   198 lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
   199 by (unfold refl_on_def) (iprover intro!: ballI)
   200 
   201 lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
   202 by (unfold refl_on_def) blast
   203 
   204 lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
   205 by (unfold refl_on_def) blast
   206 
   207 lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
   208 by (unfold refl_on_def) blast
   209 
   210 lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
   211 by (unfold refl_on_def) blast
   212 
   213 lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
   214 by (unfold refl_on_def) blast
   215 
   216 lemma refl_on_INTER:
   217   "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
   218 by (unfold refl_on_def) fast
   219 
   220 lemma refl_on_UNION:
   221   "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
   222 by (unfold refl_on_def) blast
   223 
   224 lemma refl_on_empty[simp]: "refl_on {} {}"
   225 by(simp add:refl_on_def)
   226 
   227 lemma refl_on_Id_on: "refl_on A (Id_on A)"
   228 by (rule refl_onI [OF Id_on_subset_Times Id_onI])
   229 
   230 
   231 subsection {* Antisymmetry *}
   232 
   233 lemma antisymI:
   234   "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
   235 by (unfold antisym_def) iprover
   236 
   237 lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
   238 by (unfold antisym_def) iprover
   239 
   240 lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
   241 by (unfold antisym_def) blast
   242 
   243 lemma antisym_empty [simp]: "antisym {}"
   244 by (unfold antisym_def) blast
   245 
   246 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
   247 by (unfold antisym_def) blast
   248 
   249 
   250 subsection {* Symmetry *}
   251 
   252 lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
   253 by (unfold sym_def) iprover
   254 
   255 lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
   256 by (unfold sym_def, blast)
   257 
   258 lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
   259 by (fast intro: symI dest: symD)
   260 
   261 lemma sym_Un: "sym r ==> sym s ==> sym (r \<union> s)"
   262 by (fast intro: symI dest: symD)
   263 
   264 lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)"
   265 by (fast intro: symI dest: symD)
   266 
   267 lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)"
   268 by (fast intro: symI dest: symD)
   269 
   270 lemma sym_Id_on [simp]: "sym (Id_on A)"
   271 by (rule symI) clarify
   272 
   273 
   274 subsection {* Transitivity *}
   275 
   276 lemma transI:
   277   "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
   278 by (unfold trans_def) iprover
   279 
   280 lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
   281 by (unfold trans_def) iprover
   282 
   283 lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
   284 by (fast intro: transI elim: transD)
   285 
   286 lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
   287 by (fast intro: transI elim: transD)
   288 
   289 lemma trans_Id_on [simp]: "trans (Id_on A)"
   290 by (fast intro: transI elim: transD)
   291 
   292 lemma trans_diff_Id: " trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r-Id)"
   293 unfolding antisym_def trans_def by blast
   294 
   295 subsection {* Irreflexivity *}
   296 
   297 lemma irrefl_diff_Id[simp]: "irrefl(r-Id)"
   298 by(simp add:irrefl_def)
   299 
   300 subsection {* Totality *}
   301 
   302 lemma total_on_empty[simp]: "total_on {} r"
   303 by(simp add:total_on_def)
   304 
   305 lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r"
   306 by(simp add: total_on_def)
   307 
   308 subsection {* Converse *}
   309 
   310 lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   311 by (simp add: converse_def)
   312 
   313 lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   314 by (simp add: converse_def)
   315 
   316 lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r"
   317 by (simp add: converse_def)
   318 
   319 lemma converseE [elim!]:
   320   "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P"
   321     -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
   322 by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
   323 
   324 lemma converse_converse [simp]: "(r^-1)^-1 = r"
   325 by (unfold converse_def) blast
   326 
   327 lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1"
   328 by blast
   329 
   330 lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
   331 by blast
   332 
   333 lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
   334 by blast
   335 
   336 lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
   337 by fast
   338 
   339 lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
   340 by blast
   341 
   342 lemma converse_Id [simp]: "Id^-1 = Id"
   343 by blast
   344 
   345 lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
   346 by blast
   347 
   348 lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
   349 by (unfold refl_on_def) auto
   350 
   351 lemma sym_converse [simp]: "sym (converse r) = sym r"
   352 by (unfold sym_def) blast
   353 
   354 lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
   355 by (unfold antisym_def) blast
   356 
   357 lemma trans_converse [simp]: "trans (converse r) = trans r"
   358 by (unfold trans_def) blast
   359 
   360 lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
   361 by (unfold sym_def) fast
   362 
   363 lemma sym_Un_converse: "sym (r \<union> r^-1)"
   364 by (unfold sym_def) blast
   365 
   366 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   367 by (unfold sym_def) blast
   368 
   369 lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
   370 by (auto simp: total_on_def)
   371 
   372 
   373 subsection {* Domain *}
   374 
   375 declare Domain_def [no_atp]
   376 
   377 lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
   378 by (unfold Domain_def) blast
   379 
   380 lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
   381 by (iprover intro!: iffD2 [OF Domain_iff])
   382 
   383 lemma DomainE [elim!]:
   384   "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
   385 by (iprover dest!: iffD1 [OF Domain_iff])
   386 
   387 lemma Domain_empty [simp]: "Domain {} = {}"
   388 by blast
   389 
   390 lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
   391   by auto
   392 
   393 lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
   394 by blast
   395 
   396 lemma Domain_Id [simp]: "Domain Id = UNIV"
   397 by blast
   398 
   399 lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
   400 by blast
   401 
   402 lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
   403 by blast
   404 
   405 lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
   406 by blast
   407 
   408 lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
   409 by blast
   410 
   411 lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
   412 by blast
   413 
   414 lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
   415 by(auto simp:Range_def)
   416 
   417 lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
   418 by blast
   419 
   420 lemma fst_eq_Domain: "fst ` R = Domain R"
   421 by (auto intro!:image_eqI)
   422 
   423 lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
   424 by auto
   425 
   426 lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
   427 by auto
   428 
   429 
   430 subsection {* Range *}
   431 
   432 lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
   433 by (simp add: Domain_def Range_def)
   434 
   435 lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
   436 by (unfold Range_def) (iprover intro!: converseI DomainI)
   437 
   438 lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
   439 by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
   440 
   441 lemma Range_empty [simp]: "Range {} = {}"
   442 by blast
   443 
   444 lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
   445   by auto
   446 
   447 lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
   448 by blast
   449 
   450 lemma Range_Id [simp]: "Range Id = UNIV"
   451 by blast
   452 
   453 lemma Range_Id_on [simp]: "Range (Id_on A) = A"
   454 by auto
   455 
   456 lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
   457 by blast
   458 
   459 lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
   460 by blast
   461 
   462 lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
   463 by blast
   464 
   465 lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
   466 by blast
   467 
   468 lemma Range_converse[simp]: "Range(r^-1) = Domain r"
   469 by blast
   470 
   471 lemma snd_eq_Range: "snd ` R = Range R"
   472 by (auto intro!:image_eqI)
   473 
   474 
   475 subsection {* Field *}
   476 
   477 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
   478 by(auto simp:Field_def Domain_def Range_def)
   479 
   480 lemma Field_empty[simp]: "Field {} = {}"
   481 by(auto simp:Field_def)
   482 
   483 lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
   484 by(auto simp:Field_def)
   485 
   486 lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
   487 by(auto simp:Field_def)
   488 
   489 lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
   490 by(auto simp:Field_def)
   491 
   492 lemma Field_converse[simp]: "Field(r^-1) = Field r"
   493 by(auto simp:Field_def)
   494 
   495 
   496 subsection {* Image of a set under a relation *}
   497 
   498 declare Image_def [no_atp]
   499 
   500 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
   501 by (simp add: Image_def)
   502 
   503 lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
   504 by (simp add: Image_def)
   505 
   506 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
   507 by (rule Image_iff [THEN trans]) simp
   508 
   509 lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
   510 by (unfold Image_def) blast
   511 
   512 lemma ImageE [elim!]:
   513     "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
   514 by (unfold Image_def) (iprover elim!: CollectE bexE)
   515 
   516 lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
   517   -- {* This version's more effective when we already have the required @{text a} *}
   518 by blast
   519 
   520 lemma Image_empty [simp]: "R``{} = {}"
   521 by blast
   522 
   523 lemma Image_Id [simp]: "Id `` A = A"
   524 by blast
   525 
   526 lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
   527 by blast
   528 
   529 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
   530 by blast
   531 
   532 lemma Image_Int_eq:
   533      "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
   534 by (simp add: single_valued_def, blast) 
   535 
   536 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
   537 by blast
   538 
   539 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
   540 by blast
   541 
   542 lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
   543 by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
   544 
   545 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
   546   -- {* NOT suitable for rewriting *}
   547 by blast
   548 
   549 lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
   550 by blast
   551 
   552 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
   553 by blast
   554 
   555 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
   556 by blast
   557 
   558 text{*Converse inclusion requires some assumptions*}
   559 lemma Image_INT_eq:
   560      "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
   561 apply (rule equalityI)
   562  apply (rule Image_INT_subset) 
   563 apply  (simp add: single_valued_def, blast)
   564 done
   565 
   566 lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
   567 by blast
   568 
   569 
   570 subsection {* Single valued relations *}
   571 
   572 lemma single_valuedI:
   573   "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
   574 by (unfold single_valued_def)
   575 
   576 lemma single_valuedD:
   577   "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
   578 by (simp add: single_valued_def)
   579 
   580 lemma single_valued_rel_comp:
   581   "single_valued r ==> single_valued s ==> single_valued (r O s)"
   582 by (unfold single_valued_def) blast
   583 
   584 lemma single_valued_subset:
   585   "r \<subseteq> s ==> single_valued s ==> single_valued r"
   586 by (unfold single_valued_def) blast
   587 
   588 lemma single_valued_Id [simp]: "single_valued Id"
   589 by (unfold single_valued_def) blast
   590 
   591 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
   592 by (unfold single_valued_def) blast
   593 
   594 
   595 subsection {* Graphs given by @{text Collect} *}
   596 
   597 lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}"
   598 by auto
   599 
   600 lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}"
   601 by auto
   602 
   603 lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}"
   604 by auto
   605 
   606 
   607 subsection {* Inverse image *}
   608 
   609 lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
   610 by (unfold sym_def inv_image_def) blast
   611 
   612 lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
   613   apply (unfold trans_def inv_image_def)
   614   apply (simp (no_asm))
   615   apply blast
   616   done
   617 
   618 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
   619   by (auto simp:inv_image_def)
   620 
   621 lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
   622 unfolding inv_image_def converse_def by auto
   623 
   624 
   625 subsection {* Finiteness *}
   626 
   627 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   628   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   629    apply simp
   630    apply (rule iffI)
   631     apply (erule finite_imageD [unfolded inj_on_def])
   632     apply (simp split add: split_split)
   633    apply (erule finite_imageI)
   634   apply (simp add: converse_def image_def, auto)
   635   apply (rule bexI)
   636    prefer 2 apply assumption
   637   apply simp
   638   done
   639 
   640 lemma finite_Domain: "finite r ==> finite (Domain r)"
   641   by (induct set: finite) (auto simp add: Domain_insert)
   642 
   643 lemma finite_Range: "finite r ==> finite (Range r)"
   644   by (induct set: finite) (auto simp add: Range_insert)
   645 
   646 lemma finite_Field: "finite r ==> finite (Field r)"
   647   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   648   apply (induct set: finite)
   649    apply (auto simp add: Field_def Domain_insert Range_insert)
   650   done
   651 
   652 
   653 subsection {* Miscellaneous *}
   654 
   655 text {* Version of @{thm[source] lfp_induct} for binary relations *}
   656 
   657 lemmas lfp_induct2 = 
   658   lfp_induct_set [of "(a, b)", split_format (complete)]
   659 
   660 text {* Version of @{thm[source] subsetI} for binary relations *}
   661 
   662 lemma subrelI: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
   663 by auto
   664 
   665 end