src/HOL/Relation.thy
author haftmann
Thu, 01 Mar 2012 19:34:52 +0100
changeset 47620 e9e7209eb375
parent 47567 28a01ea3523a
child 47638 807a5d219c23
permissions -rw-r--r--
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
     1 (*  Title:      HOL/Relation.thy
     2     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory; Stefan Berghofer, TU Muenchen
     3 *)
     4 
     5 header {* Relations – as sets of pairs, and binary predicates *}
     6 
     7 theory Relation
     8 imports Datatype Finite_Set
     9 begin
    10 
    11 text {* A preliminary: classical rules for reasoning on predicates *}
    12 
    13 (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *)
    14 declare predicate1D [Pure.dest?, dest?]
    15 (* CANDIDATE declare predicate1D [Pure.dest, dest] *)
    16 declare predicate2I [Pure.intro!, intro!]
    17 declare predicate2D [Pure.dest, dest]
    18 declare bot1E [elim!]
    19 declare bot2E [elim!]
    20 declare top1I [intro!]
    21 declare top2I [intro!]
    22 declare inf1I [intro!]
    23 declare inf2I [intro!]
    24 declare inf1E [elim!]
    25 declare inf2E [elim!]
    26 declare sup1I1 [intro?]
    27 declare sup2I1 [intro?]
    28 declare sup1I2 [intro?]
    29 declare sup2I2 [intro?]
    30 declare sup1E [elim!]
    31 declare sup2E [elim!]
    32 declare sup1CI [intro!]
    33 declare sup2CI [intro!]
    34 declare INF1_I [intro!]
    35 declare INF2_I [intro!]
    36 declare INF1_D [elim]
    37 declare INF2_D [elim]
    38 declare INF1_E [elim]
    39 declare INF2_E [elim]
    40 declare SUP1_I [intro]
    41 declare SUP2_I [intro]
    42 declare SUP1_E [elim!]
    43 declare SUP2_E [elim!]
    44 
    45 subsection {* Fundamental *}
    46 
    47 subsubsection {* Relations as sets of pairs *}
    48 
    49 type_synonym 'a rel = "('a * 'a) set"
    50 
    51 lemma subrelI: -- {* Version of @{thm [source] subsetI} for binary relations *}
    52   "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
    53   by auto
    54 
    55 lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
    56   "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
    57     (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
    58   using lfp_induct_set [of "(a, b)" f "prod_case P"] by auto
    59 
    60 
    61 subsubsection {* Conversions between set and predicate relations *}
    62 
    63 lemma pred_equals_eq [pred_set_conv]: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) \<longleftrightarrow> (R = S)"
    64   by (simp add: set_eq_iff fun_eq_iff)
    65 
    66 lemma pred_equals_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) = (\<lambda>x y. (x, y) \<in> S)) \<longleftrightarrow> (R = S)"
    67   by (simp add: set_eq_iff fun_eq_iff)
    68 
    69 lemma pred_subset_eq [pred_set_conv]: "((\<lambda>x. x \<in> R) \<le> (\<lambda>x. x \<in> S)) \<longleftrightarrow> (R \<subseteq> S)"
    70   by (simp add: subset_iff le_fun_def)
    71 
    72 lemma pred_subset_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) \<le> (\<lambda>x y. (x, y) \<in> S)) \<longleftrightarrow> (R \<subseteq> S)"
    73   by (simp add: subset_iff le_fun_def)
    74 
    75 lemma bot_empty_eq (* CANDIDATE [pred_set_conv] *): "\<bottom> = (\<lambda>x. x \<in> {})"
    76   by (auto simp add: fun_eq_iff)
    77 
    78 lemma bot_empty_eq2 (* CANDIDATE [pred_set_conv] *): "\<bottom> = (\<lambda>x y. (x, y) \<in> {})"
    79   by (auto simp add: fun_eq_iff)
    80 
    81 (* CANDIDATE lemma top_empty_eq [pred_set_conv]: "\<top> = (\<lambda>x. x \<in> UNIV)"
    82   by (auto simp add: fun_eq_iff) *)
    83 
    84 (* CANDIDATE lemma top_empty_eq2 [pred_set_conv]: "\<top> = (\<lambda>x y. (x, y) \<in> UNIV)"
    85   by (auto simp add: fun_eq_iff) *)
    86 
    87 lemma inf_Int_eq [pred_set_conv]: "(\<lambda>x. x \<in> R) \<sqinter> (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)"
    88   by (simp add: inf_fun_def)
    89 
    90 lemma inf_Int_eq2 [pred_set_conv]: "(\<lambda>x y. (x, y) \<in> R) \<sqinter> (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)"
    91   by (simp add: inf_fun_def)
    92 
    93 lemma sup_Un_eq [pred_set_conv]: "(\<lambda>x. x \<in> R) \<squnion> (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<union> S)"
    94   by (simp add: sup_fun_def)
    95 
    96 lemma sup_Un_eq2 [pred_set_conv]: "(\<lambda>x y. (x, y) \<in> R) \<squnion> (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<union> S)"
    97   by (simp add: sup_fun_def)
    98 
    99 lemma INF_INT_eq (* CANDIDATE [pred_set_conv] *): "(\<Sqinter>i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (\<Inter>i. r i))"
   100   by (simp add: INF_apply fun_eq_iff)
   101 
   102 lemma INF_INT_eq2 (* CANDIDATE [pred_set_conv] *): "(\<Sqinter>i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (\<Inter>i. r i))"
   103   by (simp add: INF_apply fun_eq_iff)
   104 
   105 lemma SUP_UN_eq [pred_set_conv]: "(\<Squnion>i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (\<Union>i. r i))"
   106   by (simp add: SUP_apply fun_eq_iff)
   107 
   108 lemma SUP_UN_eq2 [pred_set_conv]: "(\<Squnion>i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (\<Union>i. r i))"
   109   by (simp add: SUP_apply fun_eq_iff)
   110 
   111 
   112 subsection {* Properties of relations *}
   113 
   114 subsubsection {* Reflexivity *}
   115 
   116 definition refl_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
   117 where
   118   "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A \<and> (\<forall>x\<in>A. (x, x) \<in> r)"
   119 
   120 abbreviation refl :: "'a rel \<Rightarrow> bool"
   121 where -- {* reflexivity over a type *}
   122   "refl \<equiv> refl_on UNIV"
   123 
   124 definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
   125 where
   126   "reflp r \<longleftrightarrow> refl {(x, y). r x y}"
   127 
   128 lemma reflp_refl_eq [pred_set_conv]:
   129   "reflp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> refl r" 
   130   by (simp add: refl_on_def reflp_def)
   131 
   132 lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
   133   by (unfold refl_on_def) (iprover intro!: ballI)
   134 
   135 lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
   136   by (unfold refl_on_def) blast
   137 
   138 lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
   139   by (unfold refl_on_def) blast
   140 
   141 lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
   142   by (unfold refl_on_def) blast
   143 
   144 lemma reflpI:
   145   "(\<And>x. r x x) \<Longrightarrow> reflp r"
   146   by (auto intro: refl_onI simp add: reflp_def)
   147 
   148 lemma reflpE:
   149   assumes "reflp r"
   150   obtains "r x x"
   151   using assms by (auto dest: refl_onD simp add: reflp_def)
   152 
   153 lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
   154   by (unfold refl_on_def) blast
   155 
   156 lemma reflp_inf:
   157   "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<sqinter> s)"
   158   by (auto intro: reflpI elim: reflpE)
   159 
   160 lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
   161   by (unfold refl_on_def) blast
   162 
   163 lemma reflp_sup:
   164   "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<squnion> s)"
   165   by (auto intro: reflpI elim: reflpE)
   166 
   167 lemma refl_on_INTER:
   168   "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
   169   by (unfold refl_on_def) fast
   170 
   171 lemma refl_on_UNION:
   172   "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
   173   by (unfold refl_on_def) blast
   174 
   175 lemma refl_on_empty [simp]: "refl_on {} {}"
   176   by (simp add:refl_on_def)
   177 
   178 lemma refl_on_def' [nitpick_unfold, code]:
   179   "refl_on A r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<in> A \<and> y \<in> A) \<and> (\<forall>x \<in> A. (x, x) \<in> r)"
   180   by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
   181 
   182 
   183 subsubsection {* Irreflexivity *}
   184 
   185 definition irrefl :: "'a rel \<Rightarrow> bool"
   186 where
   187   "irrefl r \<longleftrightarrow> (\<forall>x. (x, x) \<notin> r)"
   188 
   189 lemma irrefl_distinct [code]:
   190   "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
   191   by (auto simp add: irrefl_def)
   192 
   193 
   194 subsubsection {* Symmetry *}
   195 
   196 definition sym :: "'a rel \<Rightarrow> bool"
   197 where
   198   "sym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r)"
   199 
   200 definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
   201 where
   202   "symp r \<longleftrightarrow> (\<forall>x y. r x y \<longrightarrow> r y x)"
   203 
   204 lemma symp_sym_eq [pred_set_conv]:
   205   "symp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> sym r" 
   206   by (simp add: sym_def symp_def)
   207 
   208 lemma symI:
   209   "(\<And>a b. (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r) \<Longrightarrow> sym r"
   210   by (unfold sym_def) iprover
   211 
   212 lemma sympI:
   213   "(\<And>a b. r a b \<Longrightarrow> r b a) \<Longrightarrow> symp r"
   214   by (fact symI [to_pred])
   215 
   216 lemma symE:
   217   assumes "sym r" and "(b, a) \<in> r"
   218   obtains "(a, b) \<in> r"
   219   using assms by (simp add: sym_def)
   220 
   221 lemma sympE:
   222   assumes "symp r" and "r b a"
   223   obtains "r a b"
   224   using assms by (rule symE [to_pred])
   225 
   226 lemma symD:
   227   assumes "sym r" and "(b, a) \<in> r"
   228   shows "(a, b) \<in> r"
   229   using assms by (rule symE)
   230 
   231 lemma sympD:
   232   assumes "symp r" and "r b a"
   233   shows "r a b"
   234   using assms by (rule symD [to_pred])
   235 
   236 lemma sym_Int:
   237   "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<inter> s)"
   238   by (fast intro: symI elim: symE)
   239 
   240 lemma symp_inf:
   241   "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<sqinter> s)"
   242   by (fact sym_Int [to_pred])
   243 
   244 lemma sym_Un:
   245   "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<union> s)"
   246   by (fast intro: symI elim: symE)
   247 
   248 lemma symp_sup:
   249   "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<squnion> s)"
   250   by (fact sym_Un [to_pred])
   251 
   252 lemma sym_INTER:
   253   "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (INTER S r)"
   254   by (fast intro: symI elim: symE)
   255 
   256 (* FIXME thm sym_INTER [to_pred] *)
   257 
   258 lemma sym_UNION:
   259   "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (UNION S r)"
   260   by (fast intro: symI elim: symE)
   261 
   262 (* FIXME thm sym_UNION [to_pred] *)
   263 
   264 
   265 subsubsection {* Antisymmetry *}
   266 
   267 definition antisym :: "'a rel \<Rightarrow> bool"
   268 where
   269   "antisym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r \<longrightarrow> x = y)"
   270 
   271 abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
   272 where
   273   "antisymP r \<equiv> antisym {(x, y). r x y}"
   274 
   275 lemma antisymI:
   276   "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
   277   by (unfold antisym_def) iprover
   278 
   279 lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
   280   by (unfold antisym_def) iprover
   281 
   282 lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
   283   by (unfold antisym_def) blast
   284 
   285 lemma antisym_empty [simp]: "antisym {}"
   286   by (unfold antisym_def) blast
   287 
   288 
   289 subsubsection {* Transitivity *}
   290 
   291 definition trans :: "'a rel \<Rightarrow> bool"
   292 where
   293   "trans r \<longleftrightarrow> (\<forall>x y z. (x, y) \<in> r \<longrightarrow> (y, z) \<in> r \<longrightarrow> (x, z) \<in> r)"
   294 
   295 definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
   296 where
   297   "transp r \<longleftrightarrow> (\<forall>x y z. r x y \<longrightarrow> r y z \<longrightarrow> r x z)"
   298 
   299 lemma transp_trans_eq [pred_set_conv]:
   300   "transp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> trans r" 
   301   by (simp add: trans_def transp_def)
   302 
   303 abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
   304 where -- {* FIXME drop *}
   305   "transP r \<equiv> trans {(x, y). r x y}"
   306 
   307 lemma transI:
   308   "(\<And>x y z. (x, y) \<in> r \<Longrightarrow> (y, z) \<in> r \<Longrightarrow> (x, z) \<in> r) \<Longrightarrow> trans r"
   309   by (unfold trans_def) iprover
   310 
   311 lemma transpI:
   312   "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   313   by (fact transI [to_pred])
   314 
   315 lemma transE:
   316   assumes "trans r" and "(x, y) \<in> r" and "(y, z) \<in> r"
   317   obtains "(x, z) \<in> r"
   318   using assms by (unfold trans_def) iprover
   319 
   320 lemma transpE:
   321   assumes "transp r" and "r x y" and "r y z"
   322   obtains "r x z"
   323   using assms by (rule transE [to_pred])
   324 
   325 lemma transD:
   326   assumes "trans r" and "(x, y) \<in> r" and "(y, z) \<in> r"
   327   shows "(x, z) \<in> r"
   328   using assms by (rule transE)
   329 
   330 lemma transpD:
   331   assumes "transp r" and "r x y" and "r y z"
   332   shows "r x z"
   333   using assms by (rule transD [to_pred])
   334 
   335 lemma trans_Int:
   336   "trans r \<Longrightarrow> trans s \<Longrightarrow> trans (r \<inter> s)"
   337   by (fast intro: transI elim: transE)
   338 
   339 lemma transp_inf:
   340   "transp r \<Longrightarrow> transp s \<Longrightarrow> transp (r \<sqinter> s)"
   341   by (fact trans_Int [to_pred])
   342 
   343 lemma trans_INTER:
   344   "\<forall>x\<in>S. trans (r x) \<Longrightarrow> trans (INTER S r)"
   345   by (fast intro: transI elim: transD)
   346 
   347 (* FIXME thm trans_INTER [to_pred] *)
   348 
   349 lemma trans_join [code]:
   350   "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
   351   by (auto simp add: trans_def)
   352 
   353 lemma transp_trans:
   354   "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   355   by (simp add: trans_def transp_def)
   356 
   357 
   358 subsubsection {* Totality *}
   359 
   360 definition total_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
   361 where
   362   "total_on A r \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. x \<noteq> y \<longrightarrow> (x, y) \<in> r \<or> (y, x) \<in> r)"
   363 
   364 abbreviation "total \<equiv> total_on UNIV"
   365 
   366 lemma total_on_empty [simp]: "total_on {} r"
   367   by (simp add: total_on_def)
   368 
   369 
   370 subsubsection {* Single valued relations *}
   371 
   372 definition single_valued :: "('a \<times> 'b) set \<Rightarrow> bool"
   373 where
   374   "single_valued r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z))"
   375 
   376 abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   377   "single_valuedP r \<equiv> single_valued {(x, y). r x y}"
   378 
   379 lemma single_valuedI:
   380   "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
   381   by (unfold single_valued_def)
   382 
   383 lemma single_valuedD:
   384   "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
   385   by (simp add: single_valued_def)
   386 
   387 lemma single_valued_subset:
   388   "r \<subseteq> s ==> single_valued s ==> single_valued r"
   389   by (unfold single_valued_def) blast
   390 
   391 
   392 subsection {* Relation operations *}
   393 
   394 subsubsection {* The identity relation *}
   395 
   396 definition Id :: "'a rel"
   397 where
   398   "Id = {p. \<exists>x. p = (x, x)}"
   399 
   400 lemma IdI [intro]: "(a, a) : Id"
   401   by (simp add: Id_def)
   402 
   403 lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
   404   by (unfold Id_def) (iprover elim: CollectE)
   405 
   406 lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)"
   407   by (unfold Id_def) blast
   408 
   409 lemma refl_Id: "refl Id"
   410   by (simp add: refl_on_def)
   411 
   412 lemma antisym_Id: "antisym Id"
   413   -- {* A strange result, since @{text Id} is also symmetric. *}
   414   by (simp add: antisym_def)
   415 
   416 lemma sym_Id: "sym Id"
   417   by (simp add: sym_def)
   418 
   419 lemma trans_Id: "trans Id"
   420   by (simp add: trans_def)
   421 
   422 lemma single_valued_Id [simp]: "single_valued Id"
   423   by (unfold single_valued_def) blast
   424 
   425 lemma irrefl_diff_Id [simp]: "irrefl (r - Id)"
   426   by (simp add:irrefl_def)
   427 
   428 lemma trans_diff_Id: "trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r - Id)"
   429   unfolding antisym_def trans_def by blast
   430 
   431 lemma total_on_diff_Id [simp]: "total_on A (r - Id) = total_on A r"
   432   by (simp add: total_on_def)
   433 
   434 
   435 subsubsection {* Diagonal: identity over a set *}
   436 
   437 definition Id_on  :: "'a set \<Rightarrow> 'a rel"
   438 where
   439   "Id_on A = (\<Union>x\<in>A. {(x, x)})"
   440 
   441 lemma Id_on_empty [simp]: "Id_on {} = {}"
   442   by (simp add: Id_on_def) 
   443 
   444 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   445   by (simp add: Id_on_def)
   446 
   447 lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A"
   448   by (rule Id_on_eqI) (rule refl)
   449 
   450 lemma Id_onE [elim!]:
   451   "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
   452   -- {* The general elimination rule. *}
   453   by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
   454 
   455 lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
   456   by blast
   457 
   458 lemma Id_on_def' [nitpick_unfold]:
   459   "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
   460   by auto
   461 
   462 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   463   by blast
   464 
   465 lemma refl_on_Id_on: "refl_on A (Id_on A)"
   466   by (rule refl_onI [OF Id_on_subset_Times Id_onI])
   467 
   468 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
   469   by (unfold antisym_def) blast
   470 
   471 lemma sym_Id_on [simp]: "sym (Id_on A)"
   472   by (rule symI) clarify
   473 
   474 lemma trans_Id_on [simp]: "trans (Id_on A)"
   475   by (fast intro: transI elim: transD)
   476 
   477 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
   478   by (unfold single_valued_def) blast
   479 
   480 
   481 subsubsection {* Composition *}
   482 
   483 inductive_set rel_comp  :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a \<times> 'c) set" (infixr "O" 75)
   484   for r :: "('a \<times> 'b) set" and s :: "('b \<times> 'c) set"
   485 where
   486   rel_compI [intro]: "(a, b) \<in> r \<Longrightarrow> (b, c) \<in> s \<Longrightarrow> (a, c) \<in> r O s"
   487 
   488 abbreviation pred_comp (infixr "OO" 75) where
   489   "pred_comp \<equiv> rel_compp"
   490 
   491 lemmas pred_compI = rel_compp.intros
   492 
   493 text {*
   494   For historic reasons, the elimination rules are not wholly corresponding.
   495   Feel free to consolidate this.
   496 *}
   497 
   498 inductive_cases rel_compEpair: "(a, c) \<in> r O s"
   499 inductive_cases pred_compE [elim!]: "(r OO s) a c"
   500 
   501 lemma rel_compE [elim!]: "xz \<in> r O s \<Longrightarrow>
   502   (\<And>x y z. xz = (x, z) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (y, z) \<in> s  \<Longrightarrow> P) \<Longrightarrow> P"
   503   by (cases xz) (simp, erule rel_compEpair, iprover)
   504 
   505 lemmas pred_comp_rel_comp_eq = rel_compp_rel_comp_eq
   506 
   507 lemma R_O_Id [simp]:
   508   "R O Id = R"
   509   by fast
   510 
   511 lemma Id_O_R [simp]:
   512   "Id O R = R"
   513   by fast
   514 
   515 lemma rel_comp_empty1 [simp]:
   516   "{} O R = {}"
   517   by blast
   518 
   519 (* CANDIDATE lemma pred_comp_bot1 [simp]:
   520   ""
   521   by (fact rel_comp_empty1 [to_pred]) *)
   522 
   523 lemma rel_comp_empty2 [simp]:
   524   "R O {} = {}"
   525   by blast
   526 
   527 (* CANDIDATE lemma pred_comp_bot2 [simp]:
   528   ""
   529   by (fact rel_comp_empty2 [to_pred]) *)
   530 
   531 lemma O_assoc:
   532   "(R O S) O T = R O (S O T)"
   533   by blast
   534 
   535 lemma pred_comp_assoc:
   536   "(r OO s) OO t = r OO (s OO t)"
   537   by (fact O_assoc [to_pred])
   538 
   539 lemma trans_O_subset:
   540   "trans r \<Longrightarrow> r O r \<subseteq> r"
   541   by (unfold trans_def) blast
   542 
   543 lemma transp_pred_comp_less_eq:
   544   "transp r \<Longrightarrow> r OO r \<le> r "
   545   by (fact trans_O_subset [to_pred])
   546 
   547 lemma rel_comp_mono:
   548   "r' \<subseteq> r \<Longrightarrow> s' \<subseteq> s \<Longrightarrow> r' O s' \<subseteq> r O s"
   549   by blast
   550 
   551 lemma pred_comp_mono:
   552   "r' \<le> r \<Longrightarrow> s' \<le> s \<Longrightarrow> r' OO s' \<le> r OO s "
   553   by (fact rel_comp_mono [to_pred])
   554 
   555 lemma rel_comp_subset_Sigma:
   556   "r \<subseteq> A \<times> B \<Longrightarrow> s \<subseteq> B \<times> C \<Longrightarrow> r O s \<subseteq> A \<times> C"
   557   by blast
   558 
   559 lemma rel_comp_distrib [simp]:
   560   "R O (S \<union> T) = (R O S) \<union> (R O T)" 
   561   by auto
   562 
   563 lemma pred_comp_distrib (* CANDIDATE [simp] *):
   564   "R OO (S \<squnion> T) = R OO S \<squnion> R OO T"
   565   by (fact rel_comp_distrib [to_pred])
   566 
   567 lemma rel_comp_distrib2 [simp]:
   568   "(S \<union> T) O R = (S O R) \<union> (T O R)"
   569   by auto
   570 
   571 lemma pred_comp_distrib2 (* CANDIDATE [simp] *):
   572   "(S \<squnion> T) OO R = S OO R \<squnion> T OO R"
   573   by (fact rel_comp_distrib2 [to_pred])
   574 
   575 lemma rel_comp_UNION_distrib:
   576   "s O UNION I r = (\<Union>i\<in>I. s O r i) "
   577   by auto
   578 
   579 (* FIXME thm rel_comp_UNION_distrib [to_pred] *)
   580 
   581 lemma rel_comp_UNION_distrib2:
   582   "UNION I r O s = (\<Union>i\<in>I. r i O s) "
   583   by auto
   584 
   585 (* FIXME thm rel_comp_UNION_distrib2 [to_pred] *)
   586 
   587 lemma single_valued_rel_comp:
   588   "single_valued r \<Longrightarrow> single_valued s \<Longrightarrow> single_valued (r O s)"
   589   by (unfold single_valued_def) blast
   590 
   591 lemma rel_comp_unfold:
   592   "r O s = {(x, z). \<exists>y. (x, y) \<in> r \<and> (y, z) \<in> s}"
   593   by (auto simp add: set_eq_iff)
   594 
   595 
   596 subsubsection {* Converse *}
   597 
   598 inductive_set converse :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'a) set" ("(_^-1)" [1000] 999)
   599   for r :: "('a \<times> 'b) set"
   600 where
   601   "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r^-1"
   602 
   603 notation (xsymbols)
   604   converse  ("(_\<inverse>)" [1000] 999)
   605 
   606 notation
   607   conversep ("(_^--1)" [1000] 1000)
   608 
   609 notation (xsymbols)
   610   conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
   611 
   612 lemma converseI [sym]:
   613   "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r\<inverse>"
   614   by (fact converse.intros)
   615 
   616 lemma conversepI (* CANDIDATE [sym] *):
   617   "r a b \<Longrightarrow> r\<inverse>\<inverse> b a"
   618   by (fact conversep.intros)
   619 
   620 lemma converseD [sym]:
   621   "(a, b) \<in> r\<inverse> \<Longrightarrow> (b, a) \<in> r"
   622   by (erule converse.cases) iprover
   623 
   624 lemma conversepD (* CANDIDATE [sym] *):
   625   "r\<inverse>\<inverse> b a \<Longrightarrow> r a b"
   626   by (fact converseD [to_pred])
   627 
   628 lemma converseE [elim!]:
   629   -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
   630   "yx \<in> r\<inverse> \<Longrightarrow> (\<And>x y. yx = (y, x) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> P) \<Longrightarrow> P"
   631   by (cases yx) (simp, erule converse.cases, iprover)
   632 
   633 lemmas conversepE (* CANDIDATE [elim!] *) = conversep.cases
   634 
   635 lemma converse_iff [iff]:
   636   "(a, b) \<in> r\<inverse> \<longleftrightarrow> (b, a) \<in> r"
   637   by (auto intro: converseI)
   638 
   639 lemma conversep_iff [iff]:
   640   "r\<inverse>\<inverse> a b = r b a"
   641   by (fact converse_iff [to_pred])
   642 
   643 lemma converse_converse [simp]:
   644   "(r\<inverse>)\<inverse> = r"
   645   by (simp add: set_eq_iff)
   646 
   647 lemma conversep_conversep [simp]:
   648   "(r\<inverse>\<inverse>)\<inverse>\<inverse> = r"
   649   by (fact converse_converse [to_pred])
   650 
   651 lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1"
   652   by blast
   653 
   654 lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
   655   by (iprover intro: order_antisym conversepI pred_compI
   656     elim: pred_compE dest: conversepD)
   657 
   658 lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
   659   by blast
   660 
   661 lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1"
   662   by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
   663 
   664 lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
   665   by blast
   666 
   667 lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1"
   668   by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
   669 
   670 lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
   671   by fast
   672 
   673 lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
   674   by blast
   675 
   676 lemma converse_Id [simp]: "Id^-1 = Id"
   677   by blast
   678 
   679 lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
   680   by blast
   681 
   682 lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
   683   by (unfold refl_on_def) auto
   684 
   685 lemma sym_converse [simp]: "sym (converse r) = sym r"
   686   by (unfold sym_def) blast
   687 
   688 lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
   689   by (unfold antisym_def) blast
   690 
   691 lemma trans_converse [simp]: "trans (converse r) = trans r"
   692   by (unfold trans_def) blast
   693 
   694 lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
   695   by (unfold sym_def) fast
   696 
   697 lemma sym_Un_converse: "sym (r \<union> r^-1)"
   698   by (unfold sym_def) blast
   699 
   700 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   701   by (unfold sym_def) blast
   702 
   703 lemma total_on_converse [simp]: "total_on A (r^-1) = total_on A r"
   704   by (auto simp: total_on_def)
   705 
   706 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   707   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   708    apply simp
   709    apply (rule iffI)
   710     apply (erule finite_imageD [unfolded inj_on_def])
   711     apply (simp split add: split_split)
   712    apply (erule finite_imageI)
   713   apply (simp add: set_eq_iff image_def, auto)
   714   apply (rule bexI)
   715    prefer 2 apply assumption
   716   apply simp
   717   done
   718 
   719 lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>"
   720   by (auto simp add: fun_eq_iff)
   721 
   722 lemma conversep_eq [simp]: "(op =)^--1 = op ="
   723   by (auto simp add: fun_eq_iff)
   724 
   725 lemma converse_unfold:
   726   "r\<inverse> = {(y, x). (x, y) \<in> r}"
   727   by (simp add: set_eq_iff)
   728 
   729 
   730 subsubsection {* Domain, range and field *}
   731 
   732 definition Domain :: "('a \<times> 'b) set \<Rightarrow> 'a set"
   733 where
   734   "Domain r = {x. \<exists>y. (x, y) \<in> r}"
   735 
   736 definition Range  :: "('a \<times> 'b) set \<Rightarrow> 'b set"
   737 where
   738   "Range r = Domain (r\<inverse>)"
   739 
   740 definition Field :: "'a rel \<Rightarrow> 'a set"
   741 where
   742   "Field r = Domain r \<union> Range r"
   743 
   744 declare Domain_def [no_atp]
   745 
   746 lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
   747   by (unfold Domain_def) blast
   748 
   749 lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
   750   by (iprover intro!: iffD2 [OF Domain_iff])
   751 
   752 lemma DomainE [elim!]:
   753   "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
   754   by (iprover dest!: iffD1 [OF Domain_iff])
   755 
   756 lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
   757   by (simp add: Domain_def Range_def)
   758 
   759 lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
   760   by (unfold Range_def) (iprover intro!: converseI DomainI)
   761 
   762 lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
   763   by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
   764 
   765 inductive DomainP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
   766   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
   767 where
   768   DomainPI [intro]: "r a b \<Longrightarrow> DomainP r a"
   769 
   770 inductive_cases DomainPE [elim!]: "DomainP r a"
   771 
   772 lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
   773   by (blast intro!: Orderings.order_antisym predicate1I)
   774 
   775 inductive RangeP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool"
   776   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
   777 where
   778   RangePI [intro]: "r a b \<Longrightarrow> RangeP r b"
   779 
   780 inductive_cases RangePE [elim!]: "RangeP r b"
   781 
   782 lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
   783   by (auto intro!: Orderings.order_antisym predicate1I)
   784 
   785 lemma Domain_fst [code]:
   786   "Domain r = fst ` r"
   787   by (auto simp add: image_def Bex_def)
   788 
   789 lemma Domain_empty [simp]: "Domain {} = {}"
   790   by blast
   791 
   792 lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
   793   by auto
   794 
   795 lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
   796   by blast
   797 
   798 lemma Domain_Id [simp]: "Domain Id = UNIV"
   799   by blast
   800 
   801 lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
   802   by blast
   803 
   804 lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
   805   by blast
   806 
   807 lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
   808   by blast
   809 
   810 lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
   811   by blast
   812 
   813 lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
   814   by blast
   815 
   816 lemma Domain_converse [simp]: "Domain (r\<inverse>) = Range r"
   817   by auto
   818 
   819 lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
   820   by blast
   821 
   822 lemma fst_eq_Domain: "fst ` R = Domain R"
   823   by force
   824 
   825 lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
   826   by auto
   827 
   828 lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
   829   by auto
   830 
   831 lemma Domain_Collect_split [simp]: "Domain {(x, y). P x y} = {x. EX y. P x y}"
   832   by auto
   833 
   834 lemma finite_Domain: "finite r ==> finite (Domain r)"
   835   by (induct set: finite) (auto simp add: Domain_insert)
   836 
   837 lemma Range_snd [code]:
   838   "Range r = snd ` r"
   839   by (auto simp add: image_def Bex_def)
   840 
   841 lemma Range_empty [simp]: "Range {} = {}"
   842   by blast
   843 
   844 lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
   845   by auto
   846 
   847 lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
   848   by blast
   849 
   850 lemma Range_Id [simp]: "Range Id = UNIV"
   851   by blast
   852 
   853 lemma Range_Id_on [simp]: "Range (Id_on A) = A"
   854   by auto
   855 
   856 lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
   857   by blast
   858 
   859 lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
   860   by blast
   861 
   862 lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
   863   by blast
   864 
   865 lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
   866   by blast
   867 
   868 lemma Range_converse [simp]: "Range (r\<inverse>) = Domain r"
   869   by blast
   870 
   871 lemma snd_eq_Range: "snd ` R = Range R"
   872   by force
   873 
   874 lemma Range_Collect_split [simp]: "Range {(x, y). P x y} = {y. EX x. P x y}"
   875   by auto
   876 
   877 lemma finite_Range: "finite r ==> finite (Range r)"
   878   by (induct set: finite) (auto simp add: Range_insert)
   879 
   880 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
   881   by (auto simp: Field_def Domain_def Range_def)
   882 
   883 lemma Field_empty[simp]: "Field {} = {}"
   884   by (auto simp: Field_def)
   885 
   886 lemma Field_insert [simp]: "Field (insert (a, b) r) = {a, b} \<union> Field r"
   887   by (auto simp: Field_def)
   888 
   889 lemma Field_Un [simp]: "Field (r \<union> s) = Field r \<union> Field s"
   890   by (auto simp: Field_def)
   891 
   892 lemma Field_Union [simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
   893   by (auto simp: Field_def)
   894 
   895 lemma Field_converse [simp]: "Field(r^-1) = Field r"
   896   by (auto simp: Field_def)
   897 
   898 lemma finite_Field: "finite r ==> finite (Field r)"
   899   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   900   apply (induct set: finite)
   901    apply (auto simp add: Field_def Domain_insert Range_insert)
   902   done
   903 
   904 lemma Domain_unfold:
   905   "Domain r = {x. \<exists>y. (x, y) \<in> r}"
   906   by (fact Domain_def)
   907 
   908 
   909 subsubsection {* Image of a set under a relation *}
   910 
   911 definition Image :: "('a \<times> 'b) set \<Rightarrow> 'a set \<Rightarrow> 'b set" (infixl "``" 90)
   912 where
   913   "r `` s = {y. \<exists>x\<in>s. (x, y) \<in> r}"
   914 
   915 declare Image_def [no_atp]
   916 
   917 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
   918   by (simp add: Image_def)
   919 
   920 lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
   921   by (simp add: Image_def)
   922 
   923 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
   924   by (rule Image_iff [THEN trans]) simp
   925 
   926 lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
   927   by (unfold Image_def) blast
   928 
   929 lemma ImageE [elim!]:
   930   "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
   931   by (unfold Image_def) (iprover elim!: CollectE bexE)
   932 
   933 lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
   934   -- {* This version's more effective when we already have the required @{text a} *}
   935   by blast
   936 
   937 lemma Image_empty [simp]: "R``{} = {}"
   938   by blast
   939 
   940 lemma Image_Id [simp]: "Id `` A = A"
   941   by blast
   942 
   943 lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
   944   by blast
   945 
   946 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
   947   by blast
   948 
   949 lemma Image_Int_eq:
   950      "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
   951      by (simp add: single_valued_def, blast) 
   952 
   953 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
   954   by blast
   955 
   956 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
   957   by blast
   958 
   959 lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
   960   by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
   961 
   962 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
   963   -- {* NOT suitable for rewriting *}
   964   by blast
   965 
   966 lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
   967   by blast
   968 
   969 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
   970   by blast
   971 
   972 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
   973   by blast
   974 
   975 text{*Converse inclusion requires some assumptions*}
   976 lemma Image_INT_eq:
   977      "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
   978 apply (rule equalityI)
   979  apply (rule Image_INT_subset) 
   980 apply  (simp add: single_valued_def, blast)
   981 done
   982 
   983 lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
   984   by blast
   985 
   986 lemma Image_Collect_split [simp]: "{(x, y). P x y} `` A = {y. EX x:A. P x y}"
   987   by auto
   988 
   989 
   990 subsubsection {* Inverse image *}
   991 
   992 definition inv_image :: "'b rel \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a rel"
   993 where
   994   "inv_image r f = {(x, y). (f x, f y) \<in> r}"
   995 
   996 definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
   997 where
   998   "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
   999 
  1000 lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
  1001   by (simp add: inv_image_def inv_imagep_def)
  1002 
  1003 lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
  1004   by (unfold sym_def inv_image_def) blast
  1005 
  1006 lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
  1007   apply (unfold trans_def inv_image_def)
  1008   apply (simp (no_asm))
  1009   apply blast
  1010   done
  1011 
  1012 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
  1013   by (auto simp:inv_image_def)
  1014 
  1015 lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
  1016   unfolding inv_image_def converse_unfold by auto
  1017 
  1018 lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
  1019   by (simp add: inv_imagep_def)
  1020 
  1021 
  1022 subsubsection {* Powerset *}
  1023 
  1024 definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
  1025 where
  1026   "Powp A = (\<lambda>B. \<forall>x \<in> B. A x)"
  1027 
  1028 lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)"
  1029   by (auto simp add: Powp_def fun_eq_iff)
  1030 
  1031 lemmas Powp_mono [mono] = Pow_mono [to_pred]
  1032 
  1033 end
  1034