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