src/HOL/Relation.thy
author haftmann
Sun, 26 Feb 2012 15:28:48 +0100
changeset 47547 f559866a7aa2
parent 47535 1f6c140f9c72
child 47562 72d81e789106
permissions -rw-r--r--
marked candidates for rule declarations
     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 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 
    26 subsection {* Classical rules for reasoning on predicates *}
    27 
    28 (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *)
    29 declare predicate1D [Pure.dest?, dest?]
    30 (* CANDIDATE declare predicate1D [Pure.dest, dest] *)
    31 declare predicate2I [Pure.intro!, intro!]
    32 declare predicate2D [Pure.dest, dest]
    33 declare bot1E [elim!]
    34 declare bot2E [elim!]
    35 declare top1I [intro!]
    36 declare top2I [intro!]
    37 declare inf1I [intro!]
    38 declare inf2I [intro!]
    39 declare inf1E [elim!]
    40 declare inf2E [elim!]
    41 declare sup1I1 [intro?]
    42 declare sup2I1 [intro?]
    43 declare sup1I2 [intro?]
    44 declare sup2I2 [intro?]
    45 declare sup1E [elim!]
    46 declare sup2E [elim!]
    47 declare sup1CI [intro!]
    48 declare sup2CI [intro!]
    49 declare INF1_I [intro!]
    50 declare INF2_I [intro!]
    51 declare INF1_D [elim]
    52 declare INF2_D [elim]
    53 declare INF1_E [elim]
    54 declare INF2_E [elim]
    55 declare SUP1_I [intro]
    56 declare SUP2_I [intro]
    57 declare SUP1_E [elim!]
    58 declare SUP2_E [elim!]
    59 
    60 
    61 subsection {* 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 {* Relations as sets of pairs *}
   113 
   114 type_synonym 'a rel = "('a * 'a) set"
   115 
   116 definition
   117   converse :: "('a * 'b) set => ('b * 'a) set"
   118     ("(_^-1)" [1000] 999) where
   119   "r^-1 = {(y, x). (x, y) : r}"
   120 
   121 notation (xsymbols)
   122   converse  ("(_\<inverse>)" [1000] 999)
   123 
   124 definition
   125   rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
   126     (infixr "O" 75) where
   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 
   154 definition
   155   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)"
   157 
   158 abbreviation
   159   refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
   160   "refl \<equiv> refl_on UNIV"
   161 
   162 definition
   163   sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
   164   "sym r \<longleftrightarrow> (ALL x y. (x,y): r --> (y,x): r)"
   165 
   166 definition
   167   antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
   168   "antisym r \<longleftrightarrow> (ALL x y. (x,y):r --> (y,x):r --> x=y)"
   169 
   170 definition
   171   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)"
   173 
   174 definition
   175   irrefl :: "('a * 'a) set => bool" where
   176   "irrefl r \<longleftrightarrow> (\<forall>x. (x,x) \<notin> r)"
   177 
   178 definition
   179   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)"
   181 
   182 abbreviation "total \<equiv> total_on UNIV"
   183 
   184 definition
   185   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))"
   187 
   188 definition
   189   inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
   190   "inv_image r f = {(x, y). (f x, f y) : r}"
   191 
   192 
   193 subsubsection {* The identity relation *}
   194 
   195 lemma IdI [intro]: "(a, a) : Id"
   196 by (simp add: Id_def)
   197 
   198 lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
   199 by (unfold Id_def) (iprover elim: CollectE)
   200 
   201 lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)"
   202 by (unfold Id_def) blast
   203 
   204 lemma refl_Id: "refl Id"
   205 by (simp add: refl_on_def)
   206 
   207 lemma antisym_Id: "antisym Id"
   208   -- {* A strange result, since @{text Id} is also symmetric. *}
   209 by (simp add: antisym_def)
   210 
   211 lemma sym_Id: "sym Id"
   212 by (simp add: sym_def)
   213 
   214 lemma trans_Id: "trans Id"
   215 by (simp add: trans_def)
   216 
   217 
   218 subsubsection {* Diagonal: identity over a set *}
   219 
   220 lemma Id_on_empty [simp]: "Id_on {} = {}"
   221 by (simp add: Id_on_def) 
   222 
   223 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   224 by (simp add: Id_on_def)
   225 
   226 lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A"
   227 by (rule Id_on_eqI) (rule refl)
   228 
   229 lemma Id_onE [elim!]:
   230   "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
   231   -- {* The general elimination rule. *}
   232 by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
   233 
   234 lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
   235 by blast
   236 
   237 lemma Id_on_def' [nitpick_unfold]:
   238   "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
   239 by auto
   240 
   241 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
   242 by blast
   243 
   244 
   245 subsubsection {* Composition of two relations *}
   246 
   247 lemma rel_compI [intro]:
   248   "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   249 by (unfold rel_comp_def) blast
   250 
   251 lemma rel_compE [elim!]: "xz : r O s ==>
   252   (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s  ==> P) ==> P"
   253 by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
   254 
   255 lemma rel_compEpair:
   256   "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P"
   257 by (iprover elim: rel_compE Pair_inject ssubst)
   258 
   259 lemma R_O_Id [simp]: "R O Id = R"
   260 by fast
   261 
   262 lemma Id_O_R [simp]: "Id O R = R"
   263 by fast
   264 
   265 lemma rel_comp_empty1[simp]: "{} O R = {}"
   266 by blast
   267 
   268 lemma rel_comp_empty2[simp]: "R O {} = {}"
   269 by blast
   270 
   271 lemma O_assoc: "(R O S) O T = R O (S O T)"
   272 by blast
   273 
   274 lemma trans_O_subset: "trans r ==> r O r \<subseteq> r"
   275 by (unfold trans_def) blast
   276 
   277 lemma rel_comp_mono: "r' \<subseteq> r ==> s' \<subseteq> s ==> (r' O s') \<subseteq> (r O s)"
   278 by blast
   279 
   280 lemma rel_comp_subset_Sigma:
   281     "r \<subseteq> A \<times> B ==> s \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
   282 by blast
   283 
   284 lemma rel_comp_distrib[simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)" 
   285 by auto
   286 
   287 lemma rel_comp_distrib2[simp]: "(S \<union> T) O R = (S O R) \<union> (T O R)"
   288 by auto
   289 
   290 lemma rel_comp_UNION_distrib: "s O UNION I r = UNION I (%i. s O r i)"
   291 by auto
   292 
   293 lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)"
   294 by auto
   295 
   296 
   297 subsubsection {* Reflexivity *}
   298 
   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 
   423 
   424 subsubsection {* Converse *}
   425 
   426 lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   427 by (simp add: converse_def)
   428 
   429 lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   430 by (simp add: converse_def)
   431 
   432 lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r"
   433 by (simp add: converse_def)
   434 
   435 lemma converseE [elim!]:
   436   "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P"
   437     -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
   438 by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
   439 
   440 lemma converse_converse [simp]: "(r^-1)^-1 = r"
   441 by (unfold converse_def) blast
   442 
   443 lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1"
   444 by blast
   445 
   446 lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
   447 by blast
   448 
   449 lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
   450 by blast
   451 
   452 lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
   453 by fast
   454 
   455 lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
   456 by blast
   457 
   458 lemma converse_Id [simp]: "Id^-1 = Id"
   459 by blast
   460 
   461 lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
   462 by blast
   463 
   464 lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
   465 by (unfold refl_on_def) auto
   466 
   467 lemma sym_converse [simp]: "sym (converse r) = sym r"
   468 by (unfold sym_def) blast
   469 
   470 lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
   471 by (unfold antisym_def) blast
   472 
   473 lemma trans_converse [simp]: "trans (converse r) = trans r"
   474 by (unfold trans_def) blast
   475 
   476 lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
   477 by (unfold sym_def) fast
   478 
   479 lemma sym_Un_converse: "sym (r \<union> r^-1)"
   480 by (unfold sym_def) blast
   481 
   482 lemma sym_Int_converse: "sym (r \<inter> r^-1)"
   483 by (unfold sym_def) blast
   484 
   485 lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
   486 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 
   751 lemma finite_converse [iff]: "finite (r^-1) = finite r"
   752   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
   753    apply simp
   754    apply (rule iffI)
   755     apply (erule finite_imageD [unfolded inj_on_def])
   756     apply (simp split add: split_split)
   757    apply (erule finite_imageI)
   758   apply (simp add: converse_def image_def, auto)
   759   apply (rule bexI)
   760    prefer 2 apply assumption
   761   apply simp
   762   done
   763 
   764 lemma finite_Domain: "finite r ==> finite (Domain r)"
   765   by (induct set: finite) (auto simp add: Domain_insert)
   766 
   767 lemma finite_Range: "finite r ==> finite (Range r)"
   768   by (induct set: finite) (auto simp add: Range_insert)
   769 
   770 lemma finite_Field: "finite r ==> finite (Field r)"
   771   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   772   apply (induct set: finite)
   773    apply (auto simp add: Field_def Domain_insert Range_insert)
   774   done
   775 
   776 
   777 subsubsection {* Miscellaneous *}
   778 
   779 text {* Version of @{thm[source] lfp_induct} for binary relations *}
   780 
   781 lemmas lfp_induct2 = 
   782   lfp_induct_set [of "(a, b)", split_format (complete)]
   783 
   784 text {* Version of @{thm[source] subsetI} for binary relations *}
   785 
   786 lemma subrelI: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
   787 by auto
   788 
   789 
   790 subsection {* Relations as binary predicates *}
   791 
   792 subsubsection {* Composition *}
   793 
   794 inductive pred_comp  :: "['a \<Rightarrow> 'b \<Rightarrow> bool, 'b \<Rightarrow> 'c \<Rightarrow> bool] \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> bool" (infixr "OO" 75)
   795   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and s :: "'b \<Rightarrow> 'c \<Rightarrow> bool" where
   796   pred_compI [intro]: "r a b \<Longrightarrow> s b c \<Longrightarrow> (r OO s) a c"
   797 
   798 inductive_cases pred_compE [elim!]: "(r OO s) a c"
   799 
   800 lemma pred_comp_rel_comp_eq [pred_set_conv]:
   801   "((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)"
   802   by (auto simp add: fun_eq_iff)
   803 
   804 
   805 subsubsection {* Converse *}
   806 
   807 inductive conversep :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> bool" ("(_^--1)" [1000] 1000)
   808   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   809   conversepI: "r a b \<Longrightarrow> r^--1 b a"
   810 
   811 notation (xsymbols)
   812   conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
   813 
   814 lemma conversepD:
   815   assumes ab: "r^--1 a b"
   816   shows "r b a" using ab
   817   by cases simp
   818 
   819 lemma conversep_iff [iff]: "r^--1 a b = r b a"
   820   by (iprover intro: conversepI dest: conversepD)
   821 
   822 lemma conversep_converse_eq [pred_set_conv]:
   823   "(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)"
   824   by (auto simp add: fun_eq_iff)
   825 
   826 lemma conversep_conversep [simp]: "(r^--1)^--1 = r"
   827   by (iprover intro: order_antisym conversepI dest: conversepD)
   828 
   829 lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
   830   by (iprover intro: order_antisym conversepI pred_compI
   831     elim: pred_compE dest: conversepD)
   832 
   833 lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1"
   834   by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
   835 
   836 lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1"
   837   by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
   838 
   839 lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>"
   840   by (auto simp add: fun_eq_iff)
   841 
   842 lemma conversep_eq [simp]: "(op =)^--1 = op ="
   843   by (auto simp add: fun_eq_iff)
   844 
   845 
   846 subsubsection {* Domain *}
   847 
   848 inductive DomainP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
   849   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   850   DomainPI [intro]: "r a b \<Longrightarrow> DomainP r a"
   851 
   852 inductive_cases DomainPE [elim!]: "DomainP r a"
   853 
   854 lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
   855   by (blast intro!: Orderings.order_antisym predicate1I)
   856 
   857 
   858 subsubsection {* Range *}
   859 
   860 inductive RangeP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool"
   861   for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   862   RangePI [intro]: "r a b \<Longrightarrow> RangeP r b"
   863 
   864 inductive_cases RangePE [elim!]: "RangeP r b"
   865 
   866 lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
   867   by (blast intro!: Orderings.order_antisym predicate1I)
   868 
   869 
   870 subsubsection {* Inverse image *}
   871 
   872 definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
   873   "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
   874 
   875 lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
   876   by (simp add: inv_image_def inv_imagep_def)
   877 
   878 lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
   879   by (simp add: inv_imagep_def)
   880 
   881 
   882 subsubsection {* Powerset *}
   883 
   884 definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
   885   "Powp A = (\<lambda>B. \<forall>x \<in> B. A x)"
   886 
   887 lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)"
   888   by (auto simp add: Powp_def fun_eq_iff)
   889 
   890 lemmas Powp_mono [mono] = Pow_mono [to_pred]
   891 
   892 
   893 subsubsection {* Properties of predicate relations *}
   894 
   895 abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   896   "antisymP r \<equiv> antisym {(x, y). r x y}"
   897 
   898 abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   899   "transP r \<equiv> trans {(x, y). r x y}"
   900 
   901 abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   902   "single_valuedP r \<equiv> single_valued {(x, y). r x y}"
   903 
   904 (*FIXME inconsistencies: abbreviations vs. definitions, suffix `P` vs. suffix `p`*)
   905 
   906 definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   907   "reflp r \<longleftrightarrow> refl {(x, y). r x y}"
   908 
   909 definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   910   "symp r \<longleftrightarrow> sym {(x, y). r x y}"
   911 
   912 definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   913   "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   914 
   915 lemma reflpI:
   916   "(\<And>x. r x x) \<Longrightarrow> reflp r"
   917   by (auto intro: refl_onI simp add: reflp_def)
   918 
   919 lemma reflpE:
   920   assumes "reflp r"
   921   obtains "r x x"
   922   using assms by (auto dest: refl_onD simp add: reflp_def)
   923 
   924 lemma sympI:
   925   "(\<And>x y. r x y \<Longrightarrow> r y x) \<Longrightarrow> symp r"
   926   by (auto intro: symI simp add: symp_def)
   927 
   928 lemma sympE:
   929   assumes "symp r" and "r x y"
   930   obtains "r y x"
   931   using assms by (auto dest: symD simp add: symp_def)
   932 
   933 lemma transpI:
   934   "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   935   by (auto intro: transI simp add: transp_def)
   936   
   937 lemma transpE:
   938   assumes "transp r" and "r x y" and "r y z"
   939   obtains "r x z"
   940   using assms by (auto dest: transD simp add: transp_def)
   941 
   942 no_notation
   943   bot ("\<bottom>") and
   944   top ("\<top>") and
   945   inf (infixl "\<sqinter>" 70) and
   946   sup (infixl "\<squnion>" 65) and
   947   Inf ("\<Sqinter>_" [900] 900) and
   948   Sup ("\<Squnion>_" [900] 900)
   949 
   950 no_syntax (xsymbols)
   951   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   952   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   953   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   954   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
   955 
   956 end
   957