src/HOL/Relation.thy
changeset 46692 1f8b766224f6
parent 46691 72d81e789106
child 46694 0988b22e2626
equal deleted inserted replaced
46691:72d81e789106 46692:1f8b766224f6
     5 header {* Relations – as sets of pairs, and binary predicates *}
     5 header {* Relations – as sets of pairs, and binary predicates *}
     6 
     6 
     7 theory Relation
     7 theory Relation
     8 imports Datatype Finite_Set
     8 imports Datatype Finite_Set
     9 begin
     9 begin
    10 
       
    11 notation
       
    12   bot ("\<bottom>") and
       
    13   top ("\<top>") and
       
    14   inf (infixl "\<sqinter>" 70) and
       
    15   sup (infixl "\<squnion>" 65) and
       
    16   Inf ("\<Sqinter>_" [900] 900) and
       
    17   Sup ("\<Squnion>_" [900] 900)
       
    18 
       
    19 syntax (xsymbols)
       
    20   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
       
    21   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
       
    22   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
       
    23   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
       
    24 
       
    25 
    10 
    26 subsection {* Classical rules for reasoning on predicates *}
    11 subsection {* Classical rules for reasoning on predicates *}
    27 
    12 
    28 (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *)
    13 (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *)
    29 declare predicate1D [Pure.dest?, dest?]
    14 declare predicate1D [Pure.dest?, dest?]
   111 
    96 
   112 subsection {* Relations as sets of pairs *}
    97 subsection {* Relations as sets of pairs *}
   113 
    98 
   114 type_synonym 'a rel = "('a * 'a) set"
    99 type_synonym 'a rel = "('a * 'a) set"
   115 
   100 
   116 definition
   101 lemma subrelI: -- {* Version of @{thm [source] subsetI} for binary relations *}
   117   converse :: "('a * 'b) set => ('b * 'a) set"
   102   "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
   118     ("(_^-1)" [1000] 999) where
   103   by auto
   119   "r^-1 = {(y, x). (x, y) : r}"
   104 
   120 
   105 lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
   121 notation (xsymbols)
   106   "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
   122   converse  ("(_\<inverse>)" [1000] 999)
   107     (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
   123 
   108   using lfp_induct_set [of "(a, b)" f "prod_case P"] by auto
   124 definition
   109 
   125   rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
   110 
   126     (infixr "O" 75) where
   111 subsubsection {* Reflexivity *}
   127   "r O s = {(x,z). EX y. (x, y) : r & (y, z) : s}"
       
   128 
       
   129 definition
       
   130   Image :: "[('a * 'b) set, 'a set] => 'b set"
       
   131     (infixl "``" 90) where
       
   132   "r `` s = {y. EX x:s. (x,y):r}"
       
   133 
       
   134 definition
       
   135   Id :: "('a * 'a) set" where -- {* the identity relation *}
       
   136   "Id = {p. EX x. p = (x,x)}"
       
   137 
       
   138 definition
       
   139   Id_on  :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
       
   140   "Id_on A = (\<Union>x\<in>A. {(x,x)})"
       
   141 
       
   142 definition
       
   143   Domain :: "('a * 'b) set => 'a set" where
       
   144   "Domain r = {x. EX y. (x,y):r}"
       
   145 
       
   146 definition
       
   147   Range  :: "('a * 'b) set => 'b set" where
       
   148   "Range r = Domain(r^-1)"
       
   149 
       
   150 definition
       
   151   Field :: "('a * 'a) set => 'a set" where
       
   152   "Field r = Domain r \<union> Range r"
       
   153 
   112 
   154 definition
   113 definition
   155   refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
   114   refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
   156   "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
   115   "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
   157 
   116 
   158 abbreviation
   117 abbreviation
   159   refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
   118   refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
   160   "refl \<equiv> refl_on UNIV"
   119   "refl \<equiv> refl_on UNIV"
   161 
   120 
       
   121 lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
       
   122 by (unfold refl_on_def) (iprover intro!: ballI)
       
   123 
       
   124 lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
       
   125 by (unfold refl_on_def) blast
       
   126 
       
   127 lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
       
   128 by (unfold refl_on_def) blast
       
   129 
       
   130 lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
       
   131 by (unfold refl_on_def) blast
       
   132 
       
   133 lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
       
   134 by (unfold refl_on_def) blast
       
   135 
       
   136 lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
       
   137 by (unfold refl_on_def) blast
       
   138 
       
   139 lemma refl_on_INTER:
       
   140   "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
       
   141 by (unfold refl_on_def) fast
       
   142 
       
   143 lemma refl_on_UNION:
       
   144   "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
       
   145 by (unfold refl_on_def) blast
       
   146 
       
   147 lemma refl_on_empty[simp]: "refl_on {} {}"
       
   148 by(simp add:refl_on_def)
       
   149 
       
   150 lemma refl_on_def' [nitpick_unfold, code]:
       
   151   "refl_on A r = ((\<forall>(x, y) \<in> r. x : A \<and> y : A) \<and> (\<forall>x \<in> A. (x, x) : r))"
       
   152 by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
       
   153 
       
   154 
       
   155 subsubsection {* Antisymmetry *}
       
   156 
       
   157 definition
       
   158   antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
       
   159   "antisym r \<longleftrightarrow> (ALL x y. (x,y):r --> (y,x):r --> x=y)"
       
   160 
       
   161 lemma antisymI:
       
   162   "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
       
   163 by (unfold antisym_def) iprover
       
   164 
       
   165 lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
       
   166 by (unfold antisym_def) iprover
       
   167 
       
   168 lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
       
   169 by (unfold antisym_def) blast
       
   170 
       
   171 lemma antisym_empty [simp]: "antisym {}"
       
   172 by (unfold antisym_def) blast
       
   173 
       
   174 
       
   175 subsubsection {* Symmetry *}
       
   176 
   162 definition
   177 definition
   163   sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
   178   sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
   164   "sym r \<longleftrightarrow> (ALL x y. (x,y): r --> (y,x): r)"
   179   "sym r \<longleftrightarrow> (ALL x y. (x,y): r --> (y,x): r)"
   165 
   180 
   166 definition
   181 lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
   167   antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
   182 by (unfold sym_def) iprover
   168   "antisym r \<longleftrightarrow> (ALL x y. (x,y):r --> (y,x):r --> x=y)"
   183 
       
   184 lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
       
   185 by (unfold sym_def, blast)
       
   186 
       
   187 lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
       
   188 by (fast intro: symI dest: symD)
       
   189 
       
   190 lemma sym_Un: "sym r ==> sym s ==> sym (r \<union> s)"
       
   191 by (fast intro: symI dest: symD)
       
   192 
       
   193 lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)"
       
   194 by (fast intro: symI dest: symD)
       
   195 
       
   196 lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)"
       
   197 by (fast intro: symI dest: symD)
       
   198 
       
   199 
       
   200 subsubsection {* Transitivity *}
   169 
   201 
   170 definition
   202 definition
   171   trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
   203   trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
   172   "trans r \<longleftrightarrow> (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
   204   "trans r \<longleftrightarrow> (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
   173 
   205 
       
   206 lemma trans_join [code]:
       
   207   "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
       
   208   by (auto simp add: trans_def)
       
   209 
       
   210 lemma transI:
       
   211   "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
       
   212 by (unfold trans_def) iprover
       
   213 
       
   214 lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
       
   215 by (unfold trans_def) iprover
       
   216 
       
   217 lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
       
   218 by (fast intro: transI elim: transD)
       
   219 
       
   220 lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
       
   221 by (fast intro: transI elim: transD)
       
   222 
       
   223 
       
   224 subsubsection {* Irreflexivity *}
       
   225 
   174 definition
   226 definition
   175   irrefl :: "('a * 'a) set => bool" where
   227   irrefl :: "('a * 'a) set => bool" where
   176   "irrefl r \<longleftrightarrow> (\<forall>x. (x,x) \<notin> r)"
   228   "irrefl r \<longleftrightarrow> (\<forall>x. (x,x) \<notin> r)"
   177 
   229 
       
   230 lemma irrefl_distinct [code]:
       
   231   "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
       
   232   by (auto simp add: irrefl_def)
       
   233 
       
   234 
       
   235 subsubsection {* Totality *}
       
   236 
   178 definition
   237 definition
   179   total_on :: "'a set => ('a * 'a) set => bool" where
   238   total_on :: "'a set => ('a * 'a) set => bool" where
   180   "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)"
   239   "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)"
   181 
   240 
   182 abbreviation "total \<equiv> total_on UNIV"
   241 abbreviation "total \<equiv> total_on UNIV"
   183 
   242 
       
   243 lemma total_on_empty[simp]: "total_on {} r"
       
   244 by(simp add:total_on_def)
       
   245 
       
   246 
       
   247 subsubsection {* Single valued relations *}
       
   248 
   184 definition
   249 definition
   185   single_valued :: "('a * 'b) set => bool" where
   250   single_valued :: "('a * 'b) set => bool" where
   186   "single_valued r \<longleftrightarrow> (ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z))"
   251   "single_valued r \<longleftrightarrow> (ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z))"
   187 
   252 
   188 definition
   253 lemma single_valuedI:
   189   inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
   254   "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
   190   "inv_image r f = {(x, y). (f x, f y) : r}"
   255 by (unfold single_valued_def)
       
   256 
       
   257 lemma single_valuedD:
       
   258   "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
       
   259 by (simp add: single_valued_def)
       
   260 
       
   261 lemma single_valued_subset:
       
   262   "r \<subseteq> s ==> single_valued s ==> single_valued r"
       
   263 by (unfold single_valued_def) blast
   191 
   264 
   192 
   265 
   193 subsubsection {* The identity relation *}
   266 subsubsection {* The identity relation *}
       
   267 
       
   268 definition
       
   269   Id :: "('a * 'a) set" where -- {* the identity relation *}
       
   270   "Id = {p. EX x. p = (x,x)}"
   194 
   271 
   195 lemma IdI [intro]: "(a, a) : Id"
   272 lemma IdI [intro]: "(a, a) : Id"
   196 by (simp add: Id_def)
   273 by (simp add: Id_def)
   197 
   274 
   198 lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
   275 lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
   212 by (simp add: sym_def)
   289 by (simp add: sym_def)
   213 
   290 
   214 lemma trans_Id: "trans Id"
   291 lemma trans_Id: "trans Id"
   215 by (simp add: trans_def)
   292 by (simp add: trans_def)
   216 
   293 
       
   294 lemma single_valued_Id [simp]: "single_valued Id"
       
   295   by (unfold single_valued_def) blast
       
   296 
       
   297 lemma irrefl_diff_Id [simp]: "irrefl (r - Id)"
       
   298   by (simp add:irrefl_def)
       
   299 
       
   300 lemma trans_diff_Id: "trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r - Id)"
       
   301   unfolding antisym_def trans_def by blast
       
   302 
       
   303 lemma total_on_diff_Id [simp]: "total_on A (r - Id) = total_on A r"
       
   304   by (simp add: total_on_def)
       
   305 
   217 
   306 
   218 subsubsection {* Diagonal: identity over a set *}
   307 subsubsection {* Diagonal: identity over a set *}
       
   308 
       
   309 definition
       
   310   Id_on  :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
       
   311   "Id_on A = (\<Union>x\<in>A. {(x,x)})"
   219 
   312 
   220 lemma Id_on_empty [simp]: "Id_on {} = {}"
   313 lemma Id_on_empty [simp]: "Id_on {} = {}"
   221 by (simp add: Id_on_def) 
   314 by (simp add: Id_on_def) 
   222 
   315 
   223 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   316 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   239 by auto
   332 by auto
   240 
   333 
   241 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   334 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   242 by blast
   335 by blast
   243 
   336 
       
   337 lemma refl_on_Id_on: "refl_on A (Id_on A)"
       
   338 by (rule refl_onI [OF Id_on_subset_Times Id_onI])
       
   339 
       
   340 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
       
   341 by (unfold antisym_def) blast
       
   342 
       
   343 lemma sym_Id_on [simp]: "sym (Id_on A)"
       
   344 by (rule symI) clarify
       
   345 
       
   346 lemma trans_Id_on [simp]: "trans (Id_on A)"
       
   347 by (fast intro: transI elim: transD)
       
   348 
       
   349 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
       
   350   by (unfold single_valued_def) blast
       
   351 
   244 
   352 
   245 subsubsection {* Composition of two relations *}
   353 subsubsection {* Composition of two relations *}
       
   354 
       
   355 definition
       
   356   rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
       
   357     (infixr "O" 75) where
       
   358   "r O s = {(x,z). EX y. (x, y) : r & (y, z) : s}"
   246 
   359 
   247 lemma rel_compI [intro]:
   360 lemma rel_compI [intro]:
   248   "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   361   "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   249 by (unfold rel_comp_def) blast
   362 by (unfold rel_comp_def) blast
   250 
   363 
   291 by auto
   404 by auto
   292 
   405 
   293 lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)"
   406 lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)"
   294 by auto
   407 by auto
   295 
   408 
   296 
   409 lemma single_valued_rel_comp:
   297 subsubsection {* Reflexivity *}
   410   "single_valued r ==> single_valued s ==> single_valued (r O s)"
   298 
   411 by (unfold single_valued_def) blast
   299 lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
       
   300 by (unfold refl_on_def) (iprover intro!: ballI)
       
   301 
       
   302 lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
       
   303 by (unfold refl_on_def) blast
       
   304 
       
   305 lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
       
   306 by (unfold refl_on_def) blast
       
   307 
       
   308 lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
       
   309 by (unfold refl_on_def) blast
       
   310 
       
   311 lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
       
   312 by (unfold refl_on_def) blast
       
   313 
       
   314 lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
       
   315 by (unfold refl_on_def) blast
       
   316 
       
   317 lemma refl_on_INTER:
       
   318   "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
       
   319 by (unfold refl_on_def) fast
       
   320 
       
   321 lemma refl_on_UNION:
       
   322   "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
       
   323 by (unfold refl_on_def) blast
       
   324 
       
   325 lemma refl_on_empty[simp]: "refl_on {} {}"
       
   326 by(simp add:refl_on_def)
       
   327 
       
   328 lemma refl_on_Id_on: "refl_on A (Id_on A)"
       
   329 by (rule refl_onI [OF Id_on_subset_Times Id_onI])
       
   330 
       
   331 lemma refl_on_def' [nitpick_unfold, code]:
       
   332   "refl_on A r = ((\<forall>(x, y) \<in> r. x : A \<and> y : A) \<and> (\<forall>x \<in> A. (x, x) : r))"
       
   333 by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
       
   334 
       
   335 
       
   336 subsubsection {* Antisymmetry *}
       
   337 
       
   338 lemma antisymI:
       
   339   "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
       
   340 by (unfold antisym_def) iprover
       
   341 
       
   342 lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
       
   343 by (unfold antisym_def) iprover
       
   344 
       
   345 lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
       
   346 by (unfold antisym_def) blast
       
   347 
       
   348 lemma antisym_empty [simp]: "antisym {}"
       
   349 by (unfold antisym_def) blast
       
   350 
       
   351 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
       
   352 by (unfold antisym_def) blast
       
   353 
       
   354 
       
   355 subsubsection {* Symmetry *}
       
   356 
       
   357 lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
       
   358 by (unfold sym_def) iprover
       
   359 
       
   360 lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
       
   361 by (unfold sym_def, blast)
       
   362 
       
   363 lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
       
   364 by (fast intro: symI dest: symD)
       
   365 
       
   366 lemma sym_Un: "sym r ==> sym s ==> sym (r \<union> s)"
       
   367 by (fast intro: symI dest: symD)
       
   368 
       
   369 lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)"
       
   370 by (fast intro: symI dest: symD)
       
   371 
       
   372 lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)"
       
   373 by (fast intro: symI dest: symD)
       
   374 
       
   375 lemma sym_Id_on [simp]: "sym (Id_on A)"
       
   376 by (rule symI) clarify
       
   377 
       
   378 
       
   379 subsubsection {* Transitivity *}
       
   380 
       
   381 lemma trans_join [code]:
       
   382   "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
       
   383   by (auto simp add: trans_def)
       
   384 
       
   385 lemma transI:
       
   386   "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
       
   387 by (unfold trans_def) iprover
       
   388 
       
   389 lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
       
   390 by (unfold trans_def) iprover
       
   391 
       
   392 lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
       
   393 by (fast intro: transI elim: transD)
       
   394 
       
   395 lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
       
   396 by (fast intro: transI elim: transD)
       
   397 
       
   398 lemma trans_Id_on [simp]: "trans (Id_on A)"
       
   399 by (fast intro: transI elim: transD)
       
   400 
       
   401 lemma trans_diff_Id: " trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r-Id)"
       
   402 unfolding antisym_def trans_def by blast
       
   403 
       
   404 
       
   405 subsubsection {* Irreflexivity *}
       
   406 
       
   407 lemma irrefl_distinct [code]:
       
   408   "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
       
   409   by (auto simp add: irrefl_def)
       
   410 
       
   411 lemma irrefl_diff_Id[simp]: "irrefl(r-Id)"
       
   412 by(simp add:irrefl_def)
       
   413 
       
   414 
       
   415 subsubsection {* Totality *}
       
   416 
       
   417 lemma total_on_empty[simp]: "total_on {} r"
       
   418 by(simp add:total_on_def)
       
   419 
       
   420 lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r"
       
   421 by(simp add: total_on_def)
       
   422 
   412 
   423 
   413 
   424 subsubsection {* Converse *}
   414 subsubsection {* Converse *}
       
   415 
       
   416 definition
       
   417   converse :: "('a * 'b) set => ('b * 'a) set"
       
   418     ("(_^-1)" [1000] 999) where
       
   419   "r^-1 = {(y, x). (x, y) : r}"
       
   420 
       
   421 notation (xsymbols)
       
   422   converse  ("(_\<inverse>)" [1000] 999)
   425 
   423 
   426 lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   424 lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   427 by (simp add: converse_def)
   425 by (simp add: converse_def)
   428 
   426 
   429 lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   427 lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   482 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   480 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   483 by (unfold sym_def) blast
   481 by (unfold sym_def) blast
   484 
   482 
   485 lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
   483 lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
   486 by (auto simp: total_on_def)
   484 by (auto simp: total_on_def)
   487 
       
   488 
       
   489 subsubsection {* Domain *}
       
   490 
       
   491 declare Domain_def [no_atp]
       
   492 
       
   493 lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
       
   494 by (unfold Domain_def) blast
       
   495 
       
   496 lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
       
   497 by (iprover intro!: iffD2 [OF Domain_iff])
       
   498 
       
   499 lemma DomainE [elim!]:
       
   500   "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
       
   501 by (iprover dest!: iffD1 [OF Domain_iff])
       
   502 
       
   503 lemma Domain_fst [code]:
       
   504   "Domain r = fst ` r"
       
   505   by (auto simp add: image_def Bex_def)
       
   506 
       
   507 lemma Domain_empty [simp]: "Domain {} = {}"
       
   508 by blast
       
   509 
       
   510 lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
       
   511   by auto
       
   512 
       
   513 lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
       
   514 by blast
       
   515 
       
   516 lemma Domain_Id [simp]: "Domain Id = UNIV"
       
   517 by blast
       
   518 
       
   519 lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
       
   520 by blast
       
   521 
       
   522 lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
       
   523 by blast
       
   524 
       
   525 lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
       
   526 by blast
       
   527 
       
   528 lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
       
   529 by blast
       
   530 
       
   531 lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
       
   532 by blast
       
   533 
       
   534 lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
       
   535 by(auto simp:Range_def)
       
   536 
       
   537 lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
       
   538 by blast
       
   539 
       
   540 lemma fst_eq_Domain: "fst ` R = Domain R"
       
   541   by force
       
   542 
       
   543 lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
       
   544 by auto
       
   545 
       
   546 lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
       
   547 by auto
       
   548 
       
   549 
       
   550 subsubsection {* Range *}
       
   551 
       
   552 lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
       
   553 by (simp add: Domain_def Range_def)
       
   554 
       
   555 lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
       
   556 by (unfold Range_def) (iprover intro!: converseI DomainI)
       
   557 
       
   558 lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
       
   559 by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
       
   560 
       
   561 lemma Range_snd [code]:
       
   562   "Range r = snd ` r"
       
   563   by (auto simp add: image_def Bex_def)
       
   564 
       
   565 lemma Range_empty [simp]: "Range {} = {}"
       
   566 by blast
       
   567 
       
   568 lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
       
   569   by auto
       
   570 
       
   571 lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
       
   572 by blast
       
   573 
       
   574 lemma Range_Id [simp]: "Range Id = UNIV"
       
   575 by blast
       
   576 
       
   577 lemma Range_Id_on [simp]: "Range (Id_on A) = A"
       
   578 by auto
       
   579 
       
   580 lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
       
   581 by blast
       
   582 
       
   583 lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
       
   584 by blast
       
   585 
       
   586 lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
       
   587 by blast
       
   588 
       
   589 lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
       
   590 by blast
       
   591 
       
   592 lemma Range_converse[simp]: "Range(r^-1) = Domain r"
       
   593 by blast
       
   594 
       
   595 lemma snd_eq_Range: "snd ` R = Range R"
       
   596   by force
       
   597 
       
   598 
       
   599 subsubsection {* Field *}
       
   600 
       
   601 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
       
   602 by(auto simp:Field_def Domain_def Range_def)
       
   603 
       
   604 lemma Field_empty[simp]: "Field {} = {}"
       
   605 by(auto simp:Field_def)
       
   606 
       
   607 lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
       
   608 by(auto simp:Field_def)
       
   609 
       
   610 lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
       
   611 by(auto simp:Field_def)
       
   612 
       
   613 lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
       
   614 by(auto simp:Field_def)
       
   615 
       
   616 lemma Field_converse[simp]: "Field(r^-1) = Field r"
       
   617 by(auto simp:Field_def)
       
   618 
       
   619 
       
   620 subsubsection {* Image of a set under a relation *}
       
   621 
       
   622 declare Image_def [no_atp]
       
   623 
       
   624 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
       
   625 by (simp add: Image_def)
       
   626 
       
   627 lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
       
   628 by (simp add: Image_def)
       
   629 
       
   630 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
       
   631 by (rule Image_iff [THEN trans]) simp
       
   632 
       
   633 lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
       
   634 by (unfold Image_def) blast
       
   635 
       
   636 lemma ImageE [elim!]:
       
   637     "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
       
   638 by (unfold Image_def) (iprover elim!: CollectE bexE)
       
   639 
       
   640 lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
       
   641   -- {* This version's more effective when we already have the required @{text a} *}
       
   642 by blast
       
   643 
       
   644 lemma Image_empty [simp]: "R``{} = {}"
       
   645 by blast
       
   646 
       
   647 lemma Image_Id [simp]: "Id `` A = A"
       
   648 by blast
       
   649 
       
   650 lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
       
   651 by blast
       
   652 
       
   653 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
       
   654 by blast
       
   655 
       
   656 lemma Image_Int_eq:
       
   657      "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
       
   658 by (simp add: single_valued_def, blast) 
       
   659 
       
   660 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
       
   661 by blast
       
   662 
       
   663 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
       
   664 by blast
       
   665 
       
   666 lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
       
   667 by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
       
   668 
       
   669 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
       
   670   -- {* NOT suitable for rewriting *}
       
   671 by blast
       
   672 
       
   673 lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
       
   674 by blast
       
   675 
       
   676 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
       
   677 by blast
       
   678 
       
   679 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
       
   680 by blast
       
   681 
       
   682 text{*Converse inclusion requires some assumptions*}
       
   683 lemma Image_INT_eq:
       
   684      "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
       
   685 apply (rule equalityI)
       
   686  apply (rule Image_INT_subset) 
       
   687 apply  (simp add: single_valued_def, blast)
       
   688 done
       
   689 
       
   690 lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
       
   691 by blast
       
   692 
       
   693 
       
   694 subsubsection {* Single valued relations *}
       
   695 
       
   696 lemma single_valuedI:
       
   697   "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
       
   698 by (unfold single_valued_def)
       
   699 
       
   700 lemma single_valuedD:
       
   701   "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
       
   702 by (simp add: single_valued_def)
       
   703 
       
   704 lemma single_valued_rel_comp:
       
   705   "single_valued r ==> single_valued s ==> single_valued (r O s)"
       
   706 by (unfold single_valued_def) blast
       
   707 
       
   708 lemma single_valued_subset:
       
   709   "r \<subseteq> s ==> single_valued s ==> single_valued r"
       
   710 by (unfold single_valued_def) blast
       
   711 
       
   712 lemma single_valued_Id [simp]: "single_valued Id"
       
   713 by (unfold single_valued_def) blast
       
   714 
       
   715 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
       
   716 by (unfold single_valued_def) blast
       
   717 
       
   718 
       
   719 subsubsection {* Graphs given by @{text Collect} *}
       
   720 
       
   721 lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}"
       
   722 by auto
       
   723 
       
   724 lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}"
       
   725 by auto
       
   726 
       
   727 lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}"
       
   728 by auto
       
   729 
       
   730 
       
   731 subsubsection {* Inverse image *}
       
   732 
       
   733 lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
       
   734 by (unfold sym_def inv_image_def) blast
       
   735 
       
   736 lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
       
   737   apply (unfold trans_def inv_image_def)
       
   738   apply (simp (no_asm))
       
   739   apply blast
       
   740   done
       
   741 
       
   742 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
       
   743   by (auto simp:inv_image_def)
       
   744 
       
   745 lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
       
   746 unfolding inv_image_def converse_def by auto
       
   747 
       
   748 
       
   749 subsubsection {* Finiteness *}
       
   750 
   485 
   751 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   486 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   752   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   487   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   753    apply simp
   488    apply simp
   754    apply (rule iffI)
   489    apply (rule iffI)
   759   apply (rule bexI)
   494   apply (rule bexI)
   760    prefer 2 apply assumption
   495    prefer 2 apply assumption
   761   apply simp
   496   apply simp
   762   done
   497   done
   763 
   498 
       
   499 
       
   500 subsubsection {* Domain, range and field *}
       
   501 
       
   502 definition
       
   503   Domain :: "('a * 'b) set => 'a set" where
       
   504   "Domain r = {x. EX y. (x,y):r}"
       
   505 
       
   506 definition
       
   507   Range  :: "('a * 'b) set => 'b set" where
       
   508   "Range r = Domain(r^-1)"
       
   509 
       
   510 definition
       
   511   Field :: "('a * 'a) set => 'a set" where
       
   512   "Field r = Domain r \<union> Range r"
       
   513 
       
   514 declare Domain_def [no_atp]
       
   515 
       
   516 lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
       
   517 by (unfold Domain_def) blast
       
   518 
       
   519 lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
       
   520 by (iprover intro!: iffD2 [OF Domain_iff])
       
   521 
       
   522 lemma DomainE [elim!]:
       
   523   "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
       
   524 by (iprover dest!: iffD1 [OF Domain_iff])
       
   525 
       
   526 lemma Domain_fst [code]:
       
   527   "Domain r = fst ` r"
       
   528   by (auto simp add: image_def Bex_def)
       
   529 
       
   530 lemma Domain_empty [simp]: "Domain {} = {}"
       
   531 by blast
       
   532 
       
   533 lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
       
   534   by auto
       
   535 
       
   536 lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
       
   537 by blast
       
   538 
       
   539 lemma Domain_Id [simp]: "Domain Id = UNIV"
       
   540 by blast
       
   541 
       
   542 lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
       
   543 by blast
       
   544 
       
   545 lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
       
   546 by blast
       
   547 
       
   548 lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
       
   549 by blast
       
   550 
       
   551 lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
       
   552 by blast
       
   553 
       
   554 lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
       
   555 by blast
       
   556 
       
   557 lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
       
   558 by(auto simp:Range_def)
       
   559 
       
   560 lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
       
   561 by blast
       
   562 
       
   563 lemma fst_eq_Domain: "fst ` R = Domain R"
       
   564   by force
       
   565 
       
   566 lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
       
   567 by auto
       
   568 
       
   569 lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
       
   570 by auto
       
   571 
       
   572 lemma Domain_Collect_split [simp]: "Domain {(x, y). P x y} = {x. EX y. P x y}"
       
   573 by auto
       
   574 
   764 lemma finite_Domain: "finite r ==> finite (Domain r)"
   575 lemma finite_Domain: "finite r ==> finite (Domain r)"
   765   by (induct set: finite) (auto simp add: Domain_insert)
   576   by (induct set: finite) (auto simp add: Domain_insert)
   766 
   577 
       
   578 lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
       
   579 by (simp add: Domain_def Range_def)
       
   580 
       
   581 lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
       
   582 by (unfold Range_def) (iprover intro!: converseI DomainI)
       
   583 
       
   584 lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
       
   585 by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
       
   586 
       
   587 lemma Range_snd [code]:
       
   588   "Range r = snd ` r"
       
   589   by (auto simp add: image_def Bex_def)
       
   590 
       
   591 lemma Range_empty [simp]: "Range {} = {}"
       
   592 by blast
       
   593 
       
   594 lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
       
   595   by auto
       
   596 
       
   597 lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
       
   598 by blast
       
   599 
       
   600 lemma Range_Id [simp]: "Range Id = UNIV"
       
   601 by blast
       
   602 
       
   603 lemma Range_Id_on [simp]: "Range (Id_on A) = A"
       
   604 by auto
       
   605 
       
   606 lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
       
   607 by blast
       
   608 
       
   609 lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
       
   610 by blast
       
   611 
       
   612 lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
       
   613 by blast
       
   614 
       
   615 lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
       
   616 by blast
       
   617 
       
   618 lemma Range_converse[simp]: "Range(r^-1) = Domain r"
       
   619 by blast
       
   620 
       
   621 lemma snd_eq_Range: "snd ` R = Range R"
       
   622   by force
       
   623 
       
   624 lemma Range_Collect_split [simp]: "Range {(x, y). P x y} = {y. EX x. P x y}"
       
   625 by auto
       
   626 
   767 lemma finite_Range: "finite r ==> finite (Range r)"
   627 lemma finite_Range: "finite r ==> finite (Range r)"
   768   by (induct set: finite) (auto simp add: Range_insert)
   628   by (induct set: finite) (auto simp add: Range_insert)
       
   629 
       
   630 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
       
   631 by(auto simp:Field_def Domain_def Range_def)
       
   632 
       
   633 lemma Field_empty[simp]: "Field {} = {}"
       
   634 by(auto simp:Field_def)
       
   635 
       
   636 lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
       
   637 by(auto simp:Field_def)
       
   638 
       
   639 lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
       
   640 by(auto simp:Field_def)
       
   641 
       
   642 lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
       
   643 by(auto simp:Field_def)
       
   644 
       
   645 lemma Field_converse[simp]: "Field(r^-1) = Field r"
       
   646 by(auto simp:Field_def)
   769 
   647 
   770 lemma finite_Field: "finite r ==> finite (Field r)"
   648 lemma finite_Field: "finite r ==> finite (Field r)"
   771   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   649   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   772   apply (induct set: finite)
   650   apply (induct set: finite)
   773    apply (auto simp add: Field_def Domain_insert Range_insert)
   651    apply (auto simp add: Field_def Domain_insert Range_insert)
   774   done
   652   done
   775 
   653 
   776 
   654 
   777 subsubsection {* Miscellaneous *}
   655 subsubsection {* Image of a set under a relation *}
   778 
   656 
   779 text {* Version of @{thm[source] lfp_induct} for binary relations *}
   657 definition
   780 
   658   Image :: "[('a * 'b) set, 'a set] => 'b set"
   781 lemmas lfp_induct2 = 
   659     (infixl "``" 90) where
   782   lfp_induct_set [of "(a, b)", split_format (complete)]
   660   "r `` s = {y. EX x:s. (x,y):r}"
   783 
   661 
   784 text {* Version of @{thm[source] subsetI} for binary relations *}
   662 declare Image_def [no_atp]
   785 
   663 
   786 lemma subrelI: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
   664 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
   787 by auto
   665 by (simp add: Image_def)
       
   666 
       
   667 lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
       
   668 by (simp add: Image_def)
       
   669 
       
   670 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
       
   671 by (rule Image_iff [THEN trans]) simp
       
   672 
       
   673 lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
       
   674 by (unfold Image_def) blast
       
   675 
       
   676 lemma ImageE [elim!]:
       
   677     "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
       
   678 by (unfold Image_def) (iprover elim!: CollectE bexE)
       
   679 
       
   680 lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
       
   681   -- {* This version's more effective when we already have the required @{text a} *}
       
   682 by blast
       
   683 
       
   684 lemma Image_empty [simp]: "R``{} = {}"
       
   685 by blast
       
   686 
       
   687 lemma Image_Id [simp]: "Id `` A = A"
       
   688 by blast
       
   689 
       
   690 lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
       
   691 by blast
       
   692 
       
   693 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
       
   694 by blast
       
   695 
       
   696 lemma Image_Int_eq:
       
   697      "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
       
   698 by (simp add: single_valued_def, blast) 
       
   699 
       
   700 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
       
   701 by blast
       
   702 
       
   703 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
       
   704 by blast
       
   705 
       
   706 lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
       
   707 by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
       
   708 
       
   709 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
       
   710   -- {* NOT suitable for rewriting *}
       
   711 by blast
       
   712 
       
   713 lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
       
   714 by blast
       
   715 
       
   716 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
       
   717 by blast
       
   718 
       
   719 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
       
   720 by blast
       
   721 
       
   722 text{*Converse inclusion requires some assumptions*}
       
   723 lemma Image_INT_eq:
       
   724      "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
       
   725 apply (rule equalityI)
       
   726  apply (rule Image_INT_subset) 
       
   727 apply  (simp add: single_valued_def, blast)
       
   728 done
       
   729 
       
   730 lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
       
   731 by blast
       
   732 
       
   733 lemma Image_Collect_split [simp]: "{(x, y). P x y} `` A = {y. EX x:A. P x y}"
       
   734 by auto
       
   735 
       
   736 
       
   737 subsubsection {* Inverse image *}
       
   738 
       
   739 definition
       
   740   inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
       
   741   "inv_image r f = {(x, y). (f x, f y) : r}"
       
   742 
       
   743 lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
       
   744 by (unfold sym_def inv_image_def) blast
       
   745 
       
   746 lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
       
   747   apply (unfold trans_def inv_image_def)
       
   748   apply (simp (no_asm))
       
   749   apply blast
       
   750   done
       
   751 
       
   752 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
       
   753   by (auto simp:inv_image_def)
       
   754 
       
   755 lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
       
   756 unfolding inv_image_def converse_def by auto
   788 
   757 
   789 
   758 
   790 subsection {* Relations as binary predicates *}
   759 subsection {* Relations as binary predicates *}
   791 
   760 
   792 subsubsection {* Composition *}
   761 subsubsection {* Composition *}