src/HOL/Relation.thy
author haftmann
Tue, 20 Sep 2011 21:47:52 +0200
changeset 45883 060f76635bfe
parent 45792 58eef4843641
child 46008 6e422d180de8
permissions -rw-r--r--
tuned specification and lemma distribution among theories; tuned proofs
     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_unfold, code]:
   136   "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
   137 by auto
   138 
   139 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   140 by blast
   141 
   142 
   143 subsection {* Composition of two relations *}
   144 
   145 lemma rel_compI [intro]:
   146   "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   147 by (unfold rel_comp_def) blast
   148 
   149 lemma rel_compE [elim!]: "xz : r O s ==>
   150   (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s  ==> P) ==> P"
   151 by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
   152 
   153 lemma rel_compEpair:
   154   "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P"
   155 by (iprover elim: rel_compE Pair_inject ssubst)
   156 
   157 lemma R_O_Id [simp]: "R O Id = R"
   158 by fast
   159 
   160 lemma Id_O_R [simp]: "Id O R = R"
   161 by fast
   162 
   163 lemma rel_comp_empty1[simp]: "{} O R = {}"
   164 by blast
   165 
   166 lemma rel_comp_empty2[simp]: "R O {} = {}"
   167 by blast
   168 
   169 lemma O_assoc: "(R O S) O T = R O (S O T)"
   170 by blast
   171 
   172 lemma trans_O_subset: "trans r ==> r O r \<subseteq> r"
   173 by (unfold trans_def) blast
   174 
   175 lemma rel_comp_mono: "r' \<subseteq> r ==> s' \<subseteq> s ==> (r' O s') \<subseteq> (r O s)"
   176 by blast
   177 
   178 lemma rel_comp_subset_Sigma:
   179     "r \<subseteq> A \<times> B ==> s \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
   180 by blast
   181 
   182 lemma rel_comp_distrib[simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)" 
   183 by auto
   184 
   185 lemma rel_comp_distrib2[simp]: "(S \<union> T) O R = (S O R) \<union> (T O R)"
   186 by auto
   187 
   188 lemma rel_comp_UNION_distrib: "s O UNION I r = UNION I (%i. s O r i)"
   189 by auto
   190 
   191 lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)"
   192 by auto
   193 
   194 
   195 subsection {* Reflexivity *}
   196 
   197 lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
   198 by (unfold refl_on_def) (iprover intro!: ballI)
   199 
   200 lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
   201 by (unfold refl_on_def) blast
   202 
   203 lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
   204 by (unfold refl_on_def) blast
   205 
   206 lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
   207 by (unfold refl_on_def) blast
   208 
   209 lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
   210 by (unfold refl_on_def) blast
   211 
   212 lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
   213 by (unfold refl_on_def) blast
   214 
   215 lemma refl_on_INTER:
   216   "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
   217 by (unfold refl_on_def) fast
   218 
   219 lemma refl_on_UNION:
   220   "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
   221 by (unfold refl_on_def) blast
   222 
   223 lemma refl_on_empty[simp]: "refl_on {} {}"
   224 by(simp add:refl_on_def)
   225 
   226 lemma refl_on_Id_on: "refl_on A (Id_on A)"
   227 by (rule refl_onI [OF Id_on_subset_Times Id_onI])
   228 
   229 lemma refl_on_def' [nitpick_unfold, code]:
   230   "refl_on A r = ((\<forall>(x, y) \<in> r. x : A \<and> y : A) \<and> (\<forall>x \<in> A. (x, x) : r))"
   231 by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
   232 
   233 subsection {* Antisymmetry *}
   234 
   235 lemma antisymI:
   236   "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
   237 by (unfold antisym_def) iprover
   238 
   239 lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
   240 by (unfold antisym_def) iprover
   241 
   242 lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
   243 by (unfold antisym_def) blast
   244 
   245 lemma antisym_empty [simp]: "antisym {}"
   246 by (unfold antisym_def) blast
   247 
   248 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
   249 by (unfold antisym_def) blast
   250 
   251 
   252 subsection {* Symmetry *}
   253 
   254 lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
   255 by (unfold sym_def) iprover
   256 
   257 lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
   258 by (unfold sym_def, blast)
   259 
   260 lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
   261 by (fast intro: symI dest: symD)
   262 
   263 lemma sym_Un: "sym r ==> sym s ==> sym (r \<union> s)"
   264 by (fast intro: symI dest: symD)
   265 
   266 lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)"
   267 by (fast intro: symI dest: symD)
   268 
   269 lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)"
   270 by (fast intro: symI dest: symD)
   271 
   272 lemma sym_Id_on [simp]: "sym (Id_on A)"
   273 by (rule symI) clarify
   274 
   275 
   276 subsection {* Transitivity *}
   277 
   278 lemma trans_join:
   279   "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
   280   by (auto simp add: trans_def)
   281 
   282 lemma transI:
   283   "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
   284 by (unfold trans_def) iprover
   285 
   286 lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
   287 by (unfold trans_def) iprover
   288 
   289 lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
   290 by (fast intro: transI elim: transD)
   291 
   292 lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
   293 by (fast intro: transI elim: transD)
   294 
   295 lemma trans_Id_on [simp]: "trans (Id_on A)"
   296 by (fast intro: transI elim: transD)
   297 
   298 lemma trans_diff_Id: " trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r-Id)"
   299 unfolding antisym_def trans_def by blast
   300 
   301 subsection {* Irreflexivity *}
   302 
   303 lemma irrefl_distinct:
   304   "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
   305   by (auto simp add: irrefl_def)
   306 
   307 lemma irrefl_diff_Id[simp]: "irrefl(r-Id)"
   308 by(simp add:irrefl_def)
   309 
   310 subsection {* Totality *}
   311 
   312 lemma total_on_empty[simp]: "total_on {} r"
   313 by(simp add:total_on_def)
   314 
   315 lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r"
   316 by(simp add: total_on_def)
   317 
   318 subsection {* Converse *}
   319 
   320 lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   321 by (simp add: converse_def)
   322 
   323 lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   324 by (simp add: converse_def)
   325 
   326 lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r"
   327 by (simp add: converse_def)
   328 
   329 lemma converseE [elim!]:
   330   "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P"
   331     -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
   332 by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
   333 
   334 lemma converse_converse [simp]: "(r^-1)^-1 = r"
   335 by (unfold converse_def) blast
   336 
   337 lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1"
   338 by blast
   339 
   340 lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
   341 by blast
   342 
   343 lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
   344 by blast
   345 
   346 lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
   347 by fast
   348 
   349 lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
   350 by blast
   351 
   352 lemma converse_Id [simp]: "Id^-1 = Id"
   353 by blast
   354 
   355 lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
   356 by blast
   357 
   358 lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
   359 by (unfold refl_on_def) auto
   360 
   361 lemma sym_converse [simp]: "sym (converse r) = sym r"
   362 by (unfold sym_def) blast
   363 
   364 lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
   365 by (unfold antisym_def) blast
   366 
   367 lemma trans_converse [simp]: "trans (converse r) = trans r"
   368 by (unfold trans_def) blast
   369 
   370 lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
   371 by (unfold sym_def) fast
   372 
   373 lemma sym_Un_converse: "sym (r \<union> r^-1)"
   374 by (unfold sym_def) blast
   375 
   376 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   377 by (unfold sym_def) blast
   378 
   379 lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
   380 by (auto simp: total_on_def)
   381 
   382 
   383 subsection {* Domain *}
   384 
   385 declare Domain_def [no_atp]
   386 
   387 lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
   388 by (unfold Domain_def) blast
   389 
   390 lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
   391 by (iprover intro!: iffD2 [OF Domain_iff])
   392 
   393 lemma DomainE [elim!]:
   394   "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
   395 by (iprover dest!: iffD1 [OF Domain_iff])
   396 
   397 lemma Domain_fst:
   398   "Domain r = fst ` r"
   399   by (auto simp add: image_def Bex_def)
   400 
   401 lemma Domain_empty [simp]: "Domain {} = {}"
   402 by blast
   403 
   404 lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
   405   by auto
   406 
   407 lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
   408 by blast
   409 
   410 lemma Domain_Id [simp]: "Domain Id = UNIV"
   411 by blast
   412 
   413 lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
   414 by blast
   415 
   416 lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
   417 by blast
   418 
   419 lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
   420 by blast
   421 
   422 lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
   423 by blast
   424 
   425 lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
   426 by blast
   427 
   428 lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
   429 by(auto simp:Range_def)
   430 
   431 lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
   432 by blast
   433 
   434 lemma fst_eq_Domain: "fst ` R = Domain R"
   435   by force
   436 
   437 lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
   438 by auto
   439 
   440 lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
   441 by auto
   442 
   443 
   444 subsection {* Range *}
   445 
   446 lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
   447 by (simp add: Domain_def Range_def)
   448 
   449 lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
   450 by (unfold Range_def) (iprover intro!: converseI DomainI)
   451 
   452 lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
   453 by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
   454 
   455 lemma Range_snd:
   456   "Range r = snd ` r"
   457   by (auto simp add: image_def Bex_def)
   458 
   459 lemma Range_empty [simp]: "Range {} = {}"
   460 by blast
   461 
   462 lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
   463   by auto
   464 
   465 lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
   466 by blast
   467 
   468 lemma Range_Id [simp]: "Range Id = UNIV"
   469 by blast
   470 
   471 lemma Range_Id_on [simp]: "Range (Id_on A) = A"
   472 by auto
   473 
   474 lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
   475 by blast
   476 
   477 lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
   478 by blast
   479 
   480 lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
   481 by blast
   482 
   483 lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
   484 by blast
   485 
   486 lemma Range_converse[simp]: "Range(r^-1) = Domain r"
   487 by blast
   488 
   489 lemma snd_eq_Range: "snd ` R = Range R"
   490   by force
   491 
   492 
   493 subsection {* Field *}
   494 
   495 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
   496 by(auto simp:Field_def Domain_def Range_def)
   497 
   498 lemma Field_empty[simp]: "Field {} = {}"
   499 by(auto simp:Field_def)
   500 
   501 lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
   502 by(auto simp:Field_def)
   503 
   504 lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
   505 by(auto simp:Field_def)
   506 
   507 lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
   508 by(auto simp:Field_def)
   509 
   510 lemma Field_converse[simp]: "Field(r^-1) = Field r"
   511 by(auto simp:Field_def)
   512 
   513 
   514 subsection {* Image of a set under a relation *}
   515 
   516 declare Image_def [no_atp]
   517 
   518 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
   519 by (simp add: Image_def)
   520 
   521 lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
   522 by (simp add: Image_def)
   523 
   524 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
   525 by (rule Image_iff [THEN trans]) simp
   526 
   527 lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
   528 by (unfold Image_def) blast
   529 
   530 lemma ImageE [elim!]:
   531     "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
   532 by (unfold Image_def) (iprover elim!: CollectE bexE)
   533 
   534 lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
   535   -- {* This version's more effective when we already have the required @{text a} *}
   536 by blast
   537 
   538 lemma Image_empty [simp]: "R``{} = {}"
   539 by blast
   540 
   541 lemma Image_Id [simp]: "Id `` A = A"
   542 by blast
   543 
   544 lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
   545 by blast
   546 
   547 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
   548 by blast
   549 
   550 lemma Image_Int_eq:
   551      "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
   552 by (simp add: single_valued_def, blast) 
   553 
   554 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
   555 by blast
   556 
   557 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
   558 by blast
   559 
   560 lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
   561 by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
   562 
   563 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
   564   -- {* NOT suitable for rewriting *}
   565 by blast
   566 
   567 lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
   568 by blast
   569 
   570 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
   571 by blast
   572 
   573 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
   574 by blast
   575 
   576 text{*Converse inclusion requires some assumptions*}
   577 lemma Image_INT_eq:
   578      "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
   579 apply (rule equalityI)
   580  apply (rule Image_INT_subset) 
   581 apply  (simp add: single_valued_def, blast)
   582 done
   583 
   584 lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
   585 by blast
   586 
   587 
   588 subsection {* Single valued relations *}
   589 
   590 lemma single_valuedI:
   591   "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
   592 by (unfold single_valued_def)
   593 
   594 lemma single_valuedD:
   595   "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
   596 by (simp add: single_valued_def)
   597 
   598 lemma single_valued_rel_comp:
   599   "single_valued r ==> single_valued s ==> single_valued (r O s)"
   600 by (unfold single_valued_def) blast
   601 
   602 lemma single_valued_subset:
   603   "r \<subseteq> s ==> single_valued s ==> single_valued r"
   604 by (unfold single_valued_def) blast
   605 
   606 lemma single_valued_Id [simp]: "single_valued Id"
   607 by (unfold single_valued_def) blast
   608 
   609 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
   610 by (unfold single_valued_def) blast
   611 
   612 
   613 subsection {* Graphs given by @{text Collect} *}
   614 
   615 lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}"
   616 by auto
   617 
   618 lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}"
   619 by auto
   620 
   621 lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}"
   622 by auto
   623 
   624 
   625 subsection {* Inverse image *}
   626 
   627 lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
   628 by (unfold sym_def inv_image_def) blast
   629 
   630 lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
   631   apply (unfold trans_def inv_image_def)
   632   apply (simp (no_asm))
   633   apply blast
   634   done
   635 
   636 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
   637   by (auto simp:inv_image_def)
   638 
   639 lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
   640 unfolding inv_image_def converse_def by auto
   641 
   642 
   643 subsection {* Finiteness *}
   644 
   645 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   646   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   647    apply simp
   648    apply (rule iffI)
   649     apply (erule finite_imageD [unfolded inj_on_def])
   650     apply (simp split add: split_split)
   651    apply (erule finite_imageI)
   652   apply (simp add: converse_def image_def, auto)
   653   apply (rule bexI)
   654    prefer 2 apply assumption
   655   apply simp
   656   done
   657 
   658 lemma finite_Domain: "finite r ==> finite (Domain r)"
   659   by (induct set: finite) (auto simp add: Domain_insert)
   660 
   661 lemma finite_Range: "finite r ==> finite (Range r)"
   662   by (induct set: finite) (auto simp add: Range_insert)
   663 
   664 lemma finite_Field: "finite r ==> finite (Field r)"
   665   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   666   apply (induct set: finite)
   667    apply (auto simp add: Field_def Domain_insert Range_insert)
   668   done
   669 
   670 
   671 subsection {* Miscellaneous *}
   672 
   673 text {* Version of @{thm[source] lfp_induct} for binary relations *}
   674 
   675 lemmas lfp_induct2 = 
   676   lfp_induct_set [of "(a, b)", split_format (complete)]
   677 
   678 text {* Version of @{thm[source] subsetI} for binary relations *}
   679 
   680 lemma subrelI: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
   681 by auto
   682 
   683 end