tuned structure
authorhaftmann
Sun, 26 Feb 2012 21:26:28 +0100
changeset 475650988b22e2626
parent 47564 78bada13da46
child 47566 b779c3f21f05
tuned structure
src/HOL/Relation.thy
     1.1 --- a/src/HOL/Relation.thy	Sun Feb 26 21:25:54 2012 +0100
     1.2 +++ b/src/HOL/Relation.thy	Sun Feb 26 21:26:28 2012 +0100
     1.3 @@ -8,7 +8,7 @@
     1.4  imports Datatype Finite_Set
     1.5  begin
     1.6  
     1.7 -subsection {* Classical rules for reasoning on predicates *}
     1.8 +text {* A preliminary: classical rules for reasoning on predicates *}
     1.9  
    1.10  (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *)
    1.11  declare predicate1D [Pure.dest?, dest?]
    1.12 @@ -42,8 +42,23 @@
    1.13  declare SUP1_E [elim!]
    1.14  declare SUP2_E [elim!]
    1.15  
    1.16 +subsection {* Fundamental *}
    1.17  
    1.18 -subsection {* Conversions between set and predicate relations *}
    1.19 +subsubsection {* Relations as sets of pairs *}
    1.20 +
    1.21 +type_synonym 'a rel = "('a * 'a) set"
    1.22 +
    1.23 +lemma subrelI: -- {* Version of @{thm [source] subsetI} for binary relations *}
    1.24 +  "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
    1.25 +  by auto
    1.26 +
    1.27 +lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
    1.28 +  "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
    1.29 +    (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
    1.30 +  using lfp_induct_set [of "(a, b)" f "prod_case P"] by auto
    1.31 +
    1.32 +
    1.33 +subsubsection {* Conversions between set and predicate relations *}
    1.34  
    1.35  lemma pred_equals_eq [pred_set_conv]: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) \<longleftrightarrow> (R = S)"
    1.36    by (simp add: set_eq_iff fun_eq_iff)
    1.37 @@ -94,30 +109,21 @@
    1.38    by (simp add: SUP_apply fun_eq_iff)
    1.39  
    1.40  
    1.41 -subsection {* Relations as sets of pairs *}
    1.42 -
    1.43 -type_synonym 'a rel = "('a * 'a) set"
    1.44 -
    1.45 -lemma subrelI: -- {* Version of @{thm [source] subsetI} for binary relations *}
    1.46 -  "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
    1.47 -  by auto
    1.48 -
    1.49 -lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
    1.50 -  "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
    1.51 -    (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
    1.52 -  using lfp_induct_set [of "(a, b)" f "prod_case P"] by auto
    1.53 -
    1.54 +subsection {* Properties of relations *}
    1.55  
    1.56  subsubsection {* Reflexivity *}
    1.57  
    1.58  definition
    1.59 -  refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
    1.60 +  refl_on :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where
    1.61    "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
    1.62  
    1.63  abbreviation
    1.64    refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
    1.65    "refl \<equiv> refl_on UNIV"
    1.66  
    1.67 +definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
    1.68 +  "reflp r \<longleftrightarrow> refl {(x, y). r x y}"
    1.69 +
    1.70  lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
    1.71  by (unfold refl_on_def) (iprover intro!: ballI)
    1.72  
    1.73 @@ -130,6 +136,15 @@
    1.74  lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
    1.75  by (unfold refl_on_def) blast
    1.76  
    1.77 +lemma reflpI:
    1.78 +  "(\<And>x. r x x) \<Longrightarrow> reflp r"
    1.79 +  by (auto intro: refl_onI simp add: reflp_def)
    1.80 +
    1.81 +lemma reflpE:
    1.82 +  assumes "reflp r"
    1.83 +  obtains "r x x"
    1.84 +  using assms by (auto dest: refl_onD simp add: reflp_def)
    1.85 +
    1.86  lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
    1.87  by (unfold refl_on_def) blast
    1.88  
    1.89 @@ -152,30 +167,21 @@
    1.90  by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
    1.91  
    1.92  
    1.93 -subsubsection {* Antisymmetry *}
    1.94 +subsubsection {* Irreflexivity *}
    1.95  
    1.96  definition
    1.97 -  antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
    1.98 -  "antisym r \<longleftrightarrow> (ALL x y. (x,y):r --> (y,x):r --> x=y)"
    1.99 +  irrefl :: "('a * 'a) set => bool" where
   1.100 +  "irrefl r \<longleftrightarrow> (\<forall>x. (x,x) \<notin> r)"
   1.101  
   1.102 -lemma antisymI:
   1.103 -  "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
   1.104 -by (unfold antisym_def) iprover
   1.105 -
   1.106 -lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
   1.107 -by (unfold antisym_def) iprover
   1.108 -
   1.109 -lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
   1.110 -by (unfold antisym_def) blast
   1.111 -
   1.112 -lemma antisym_empty [simp]: "antisym {}"
   1.113 -by (unfold antisym_def) blast
   1.114 +lemma irrefl_distinct [code]:
   1.115 +  "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
   1.116 +  by (auto simp add: irrefl_def)
   1.117  
   1.118  
   1.119  subsubsection {* Symmetry *}
   1.120  
   1.121  definition
   1.122 -  sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
   1.123 +  sym :: "('a * 'a) set => bool" where
   1.124    "sym r \<longleftrightarrow> (ALL x y. (x,y): r --> (y,x): r)"
   1.125  
   1.126  lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
   1.127 @@ -184,6 +190,18 @@
   1.128  lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
   1.129  by (unfold sym_def, blast)
   1.130  
   1.131 +definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.132 +  "symp r \<longleftrightarrow> sym {(x, y). r x y}"
   1.133 +
   1.134 +lemma sympI:
   1.135 +  "(\<And>x y. r x y \<Longrightarrow> r y x) \<Longrightarrow> symp r"
   1.136 +  by (auto intro: symI simp add: symp_def)
   1.137 +
   1.138 +lemma sympE:
   1.139 +  assumes "symp r" and "r x y"
   1.140 +  obtains "r y x"
   1.141 +  using assms by (auto dest: symD simp add: symp_def)
   1.142 +
   1.143  lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
   1.144  by (fast intro: symI dest: symD)
   1.145  
   1.146 @@ -197,16 +215,35 @@
   1.147  by (fast intro: symI dest: symD)
   1.148  
   1.149  
   1.150 +subsubsection {* Antisymmetry *}
   1.151 +
   1.152 +definition
   1.153 +  antisym :: "('a * 'a) set => bool" where
   1.154 +  "antisym r \<longleftrightarrow> (ALL x y. (x,y):r --> (y,x):r --> x=y)"
   1.155 +
   1.156 +lemma antisymI:
   1.157 +  "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
   1.158 +by (unfold antisym_def) iprover
   1.159 +
   1.160 +lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
   1.161 +by (unfold antisym_def) iprover
   1.162 +
   1.163 +abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.164 +  "antisymP r \<equiv> antisym {(x, y). r x y}"
   1.165 +
   1.166 +lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
   1.167 +by (unfold antisym_def) blast
   1.168 +
   1.169 +lemma antisym_empty [simp]: "antisym {}"
   1.170 +by (unfold antisym_def) blast
   1.171 +
   1.172 +
   1.173  subsubsection {* Transitivity *}
   1.174  
   1.175  definition
   1.176 -  trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
   1.177 +  trans :: "('a * 'a) set => bool" where
   1.178    "trans r \<longleftrightarrow> (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
   1.179  
   1.180 -lemma trans_join [code]:
   1.181 -  "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
   1.182 -  by (auto simp add: trans_def)
   1.183 -
   1.184  lemma transI:
   1.185    "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
   1.186  by (unfold trans_def) iprover
   1.187 @@ -214,22 +251,30 @@
   1.188  lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
   1.189  by (unfold trans_def) iprover
   1.190  
   1.191 +abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.192 +  "transP r \<equiv> trans {(x, y). r x y}"
   1.193 +
   1.194 +definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.195 +  "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   1.196 +
   1.197 +lemma transpI:
   1.198 +  "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   1.199 +  by (auto intro: transI simp add: transp_def)
   1.200 +  
   1.201 +lemma transpE:
   1.202 +  assumes "transp r" and "r x y" and "r y z"
   1.203 +  obtains "r x z"
   1.204 +  using assms by (auto dest: transD simp add: transp_def)
   1.205 +
   1.206  lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
   1.207  by (fast intro: transI elim: transD)
   1.208  
   1.209  lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
   1.210  by (fast intro: transI elim: transD)
   1.211  
   1.212 -
   1.213 -subsubsection {* Irreflexivity *}
   1.214 -
   1.215 -definition
   1.216 -  irrefl :: "('a * 'a) set => bool" where
   1.217 -  "irrefl r \<longleftrightarrow> (\<forall>x. (x,x) \<notin> r)"
   1.218 -
   1.219 -lemma irrefl_distinct [code]:
   1.220 -  "irrefl r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<noteq> y)"
   1.221 -  by (auto simp add: irrefl_def)
   1.222 +lemma trans_join [code]:
   1.223 +  "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
   1.224 +  by (auto simp add: trans_def)
   1.225  
   1.226  
   1.227  subsubsection {* Totality *}
   1.228 @@ -258,15 +303,20 @@
   1.229    "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
   1.230  by (simp add: single_valued_def)
   1.231  
   1.232 +abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   1.233 +  "single_valuedP r \<equiv> single_valued {(x, y). r x y}"
   1.234 +
   1.235  lemma single_valued_subset:
   1.236    "r \<subseteq> s ==> single_valued s ==> single_valued r"
   1.237  by (unfold single_valued_def) blast
   1.238  
   1.239  
   1.240 +subsection {* Relation operations *}
   1.241 +
   1.242  subsubsection {* The identity relation *}
   1.243  
   1.244  definition
   1.245 -  Id :: "('a * 'a) set" where -- {* the identity relation *}
   1.246 +  Id :: "('a * 'a) set" where
   1.247    "Id = {p. EX x. p = (x,x)}"
   1.248  
   1.249  lemma IdI [intro]: "(a, a) : Id"
   1.250 @@ -307,7 +357,7 @@
   1.251  subsubsection {* Diagonal: identity over a set *}
   1.252  
   1.253  definition
   1.254 -  Id_on  :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
   1.255 +  Id_on  :: "'a set => ('a * 'a) set" where
   1.256    "Id_on A = (\<Union>x\<in>A. {(x,x)})"
   1.257  
   1.258  lemma Id_on_empty [simp]: "Id_on {} = {}"
   1.259 @@ -350,12 +400,11 @@
   1.260    by (unfold single_valued_def) blast
   1.261  
   1.262  
   1.263 -subsubsection {* Composition of two relations *}
   1.264 +subsubsection {* Composition *}
   1.265  
   1.266 -definition
   1.267 -  rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
   1.268 -    (infixr "O" 75) where
   1.269 -  "r O s = {(x,z). EX y. (x, y) : r & (y, z) : s}"
   1.270 +definition rel_comp  :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a * 'c) set" (infixr "O" 75)
   1.271 +where
   1.272 +  "r O s = {(x, z). \<exists>y. (x, y) \<in> r \<and> (y, z) \<in> s}"
   1.273  
   1.274  lemma rel_compI [intro]:
   1.275    "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
   1.276 @@ -365,6 +414,17 @@
   1.277    (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s  ==> P) ==> P"
   1.278  by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
   1.279  
   1.280 +inductive pred_comp :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> bool" (infixr "OO" 75)
   1.281 +for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and s :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
   1.282 +where
   1.283 +  pred_compI [intro]: "r a b \<Longrightarrow> s b c \<Longrightarrow> (r OO s) a c"
   1.284 +
   1.285 +inductive_cases pred_compE [elim!]: "(r OO s) a c"
   1.286 +
   1.287 +lemma pred_comp_rel_comp_eq [pred_set_conv]:
   1.288 +  "((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)"
   1.289 +  by (auto simp add: fun_eq_iff)
   1.290 +
   1.291  lemma rel_compEpair:
   1.292    "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P"
   1.293  by (iprover elim: rel_compE Pair_inject ssubst)
   1.294 @@ -421,19 +481,58 @@
   1.295  notation (xsymbols)
   1.296    converse  ("(_\<inverse>)" [1000] 999)
   1.297  
   1.298 -lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   1.299 -by (simp add: converse_def)
   1.300 +lemma converseI [sym]: "(a, b) : r ==> (b, a) : r^-1"
   1.301 +  by (simp add: converse_def)
   1.302  
   1.303 -lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
   1.304 -by (simp add: converse_def)
   1.305 -
   1.306 -lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r"
   1.307 -by (simp add: converse_def)
   1.308 +lemma converseD [sym]: "(a,b) : r^-1 ==> (b, a) : r"
   1.309 +  by (simp add: converse_def)
   1.310  
   1.311  lemma converseE [elim!]:
   1.312    "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P"
   1.313      -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
   1.314 -by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
   1.315 +  by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
   1.316 +
   1.317 +lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
   1.318 +  by (simp add: converse_def)
   1.319 +
   1.320 +inductive conversep :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> bool" ("(_^--1)" [1000] 1000)
   1.321 +  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.322 +  conversepI: "r a b \<Longrightarrow> r^--1 b a"
   1.323 +
   1.324 +notation (xsymbols)
   1.325 +  conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
   1.326 +
   1.327 +lemma conversepD:
   1.328 +  assumes ab: "r^--1 a b"
   1.329 +  shows "r b a" using ab
   1.330 +  by cases simp
   1.331 +
   1.332 +lemma conversep_iff [iff]: "r^--1 a b = r b a"
   1.333 +  by (iprover intro: conversepI dest: conversepD)
   1.334 +
   1.335 +lemma conversep_converse_eq [pred_set_conv]:
   1.336 +  "(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)"
   1.337 +  apply (auto simp add: fun_eq_iff)
   1.338 +  oops
   1.339 +
   1.340 +lemma conversep_conversep [simp]: "(r^--1)^--1 = r"
   1.341 +  by (iprover intro: order_antisym conversepI dest: conversepD)
   1.342 +
   1.343 +lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
   1.344 +  by (iprover intro: order_antisym conversepI pred_compI
   1.345 +    elim: pred_compE dest: conversepD)
   1.346 +
   1.347 +lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1"
   1.348 +  by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
   1.349 +
   1.350 +lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1"
   1.351 +  by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
   1.352 +
   1.353 +lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>"
   1.354 +  by (auto simp add: fun_eq_iff)
   1.355 +
   1.356 +lemma conversep_eq [simp]: "(op =)^--1 = op ="
   1.357 +  by (auto simp add: fun_eq_iff)
   1.358  
   1.359  lemma converse_converse [simp]: "(r^-1)^-1 = r"
   1.360  by (unfold converse_def) blast
   1.361 @@ -523,58 +622,6 @@
   1.362    "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
   1.363  by (iprover dest!: iffD1 [OF Domain_iff])
   1.364  
   1.365 -lemma Domain_fst [code]:
   1.366 -  "Domain r = fst ` r"
   1.367 -  by (auto simp add: image_def Bex_def)
   1.368 -
   1.369 -lemma Domain_empty [simp]: "Domain {} = {}"
   1.370 -by blast
   1.371 -
   1.372 -lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
   1.373 -  by auto
   1.374 -
   1.375 -lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
   1.376 -by blast
   1.377 -
   1.378 -lemma Domain_Id [simp]: "Domain Id = UNIV"
   1.379 -by blast
   1.380 -
   1.381 -lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
   1.382 -by blast
   1.383 -
   1.384 -lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
   1.385 -by blast
   1.386 -
   1.387 -lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
   1.388 -by blast
   1.389 -
   1.390 -lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
   1.391 -by blast
   1.392 -
   1.393 -lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
   1.394 -by blast
   1.395 -
   1.396 -lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
   1.397 -by(auto simp:Range_def)
   1.398 -
   1.399 -lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
   1.400 -by blast
   1.401 -
   1.402 -lemma fst_eq_Domain: "fst ` R = Domain R"
   1.403 -  by force
   1.404 -
   1.405 -lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
   1.406 -by auto
   1.407 -
   1.408 -lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
   1.409 -by auto
   1.410 -
   1.411 -lemma Domain_Collect_split [simp]: "Domain {(x, y). P x y} = {x. EX y. P x y}"
   1.412 -by auto
   1.413 -
   1.414 -lemma finite_Domain: "finite r ==> finite (Domain r)"
   1.415 -  by (induct set: finite) (auto simp add: Domain_insert)
   1.416 -
   1.417  lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
   1.418  by (simp add: Domain_def Range_def)
   1.419  
   1.420 @@ -584,66 +631,136 @@
   1.421  lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
   1.422  by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
   1.423  
   1.424 +inductive DomainP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
   1.425 +  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.426 +  DomainPI [intro]: "r a b \<Longrightarrow> DomainP r a"
   1.427 +
   1.428 +inductive_cases DomainPE [elim!]: "DomainP r a"
   1.429 +
   1.430 +lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
   1.431 +  by (blast intro!: Orderings.order_antisym predicate1I)
   1.432 +
   1.433 +inductive RangeP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool"
   1.434 +  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.435 +  RangePI [intro]: "r a b \<Longrightarrow> RangeP r b"
   1.436 +
   1.437 +inductive_cases RangePE [elim!]: "RangeP r b"
   1.438 +
   1.439 +lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
   1.440 +  by (auto intro!: Orderings.order_antisym predicate1I)
   1.441 +
   1.442 +lemma Domain_fst [code]:
   1.443 +  "Domain r = fst ` r"
   1.444 +  by (auto simp add: image_def Bex_def)
   1.445 +
   1.446 +lemma Domain_empty [simp]: "Domain {} = {}"
   1.447 +  by blast
   1.448 +
   1.449 +lemma Domain_empty_iff: "Domain r = {} \<longleftrightarrow> r = {}"
   1.450 +  by auto
   1.451 +
   1.452 +lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
   1.453 +  by blast
   1.454 +
   1.455 +lemma Domain_Id [simp]: "Domain Id = UNIV"
   1.456 +  by blast
   1.457 +
   1.458 +lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
   1.459 +  by blast
   1.460 +
   1.461 +lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
   1.462 +  by blast
   1.463 +
   1.464 +lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
   1.465 +  by blast
   1.466 +
   1.467 +lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
   1.468 +  by blast
   1.469 +
   1.470 +lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
   1.471 +  by blast
   1.472 +
   1.473 +lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
   1.474 +  by(auto simp: Range_def)
   1.475 +
   1.476 +lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
   1.477 +  by blast
   1.478 +
   1.479 +lemma fst_eq_Domain: "fst ` R = Domain R"
   1.480 +  by force
   1.481 +
   1.482 +lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
   1.483 +  by auto
   1.484 +
   1.485 +lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
   1.486 +  by auto
   1.487 +
   1.488 +lemma Domain_Collect_split [simp]: "Domain {(x, y). P x y} = {x. EX y. P x y}"
   1.489 +  by auto
   1.490 +
   1.491 +lemma finite_Domain: "finite r ==> finite (Domain r)"
   1.492 +  by (induct set: finite) (auto simp add: Domain_insert)
   1.493 +
   1.494  lemma Range_snd [code]:
   1.495    "Range r = snd ` r"
   1.496    by (auto simp add: image_def Bex_def)
   1.497  
   1.498  lemma Range_empty [simp]: "Range {} = {}"
   1.499 -by blast
   1.500 +  by blast
   1.501  
   1.502  lemma Range_empty_iff: "Range r = {} \<longleftrightarrow> r = {}"
   1.503    by auto
   1.504  
   1.505  lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
   1.506 -by blast
   1.507 +  by blast
   1.508  
   1.509  lemma Range_Id [simp]: "Range Id = UNIV"
   1.510 -by blast
   1.511 +  by blast
   1.512  
   1.513  lemma Range_Id_on [simp]: "Range (Id_on A) = A"
   1.514 -by auto
   1.515 +  by auto
   1.516  
   1.517  lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
   1.518 -by blast
   1.519 +  by blast
   1.520  
   1.521  lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
   1.522 -by blast
   1.523 +  by blast
   1.524  
   1.525  lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
   1.526 -by blast
   1.527 +  by blast
   1.528  
   1.529  lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
   1.530 -by blast
   1.531 +  by blast
   1.532  
   1.533 -lemma Range_converse[simp]: "Range(r^-1) = Domain r"
   1.534 -by blast
   1.535 +lemma Range_converse [simp]: "Range(r^-1) = Domain r"
   1.536 +  by blast
   1.537  
   1.538  lemma snd_eq_Range: "snd ` R = Range R"
   1.539    by force
   1.540  
   1.541  lemma Range_Collect_split [simp]: "Range {(x, y). P x y} = {y. EX x. P x y}"
   1.542 -by auto
   1.543 +  by auto
   1.544  
   1.545  lemma finite_Range: "finite r ==> finite (Range r)"
   1.546    by (induct set: finite) (auto simp add: Range_insert)
   1.547  
   1.548  lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
   1.549 -by(auto simp:Field_def Domain_def Range_def)
   1.550 +  by (auto simp: Field_def Domain_def Range_def)
   1.551  
   1.552  lemma Field_empty[simp]: "Field {} = {}"
   1.553 -by(auto simp:Field_def)
   1.554 +  by (auto simp: Field_def)
   1.555  
   1.556 -lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
   1.557 -by(auto simp:Field_def)
   1.558 +lemma Field_insert [simp]: "Field (insert (a, b) r) = {a, b} \<union> Field r"
   1.559 +  by (auto simp: Field_def)
   1.560  
   1.561 -lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
   1.562 -by(auto simp:Field_def)
   1.563 +lemma Field_Un [simp]: "Field (r \<union> s) = Field r \<union> Field s"
   1.564 +  by (auto simp: Field_def)
   1.565  
   1.566 -lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
   1.567 -by(auto simp:Field_def)
   1.568 +lemma Field_Union [simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
   1.569 +  by (auto simp: Field_def)
   1.570  
   1.571 -lemma Field_converse[simp]: "Field(r^-1) = Field r"
   1.572 -by(auto simp:Field_def)
   1.573 +lemma Field_converse [simp]: "Field(r^-1) = Field r"
   1.574 +  by (auto simp: Field_def)
   1.575  
   1.576  lemma finite_Field: "finite r ==> finite (Field r)"
   1.577    -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
   1.578 @@ -740,6 +857,12 @@
   1.579    inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
   1.580    "inv_image r f = {(x, y). (f x, f y) : r}"
   1.581  
   1.582 +definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
   1.583 +  "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
   1.584 +
   1.585 +lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
   1.586 +  by (simp add: inv_image_def inv_imagep_def)
   1.587 +
   1.588  lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
   1.589  by (unfold sym_def inv_image_def) blast
   1.590  
   1.591 @@ -755,95 +878,6 @@
   1.592  lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
   1.593  unfolding inv_image_def converse_def by auto
   1.594  
   1.595 -
   1.596 -subsection {* Relations as binary predicates *}
   1.597 -
   1.598 -subsubsection {* Composition *}
   1.599 -
   1.600 -inductive pred_comp  :: "['a \<Rightarrow> 'b \<Rightarrow> bool, 'b \<Rightarrow> 'c \<Rightarrow> bool] \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> bool" (infixr "OO" 75)
   1.601 -  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and s :: "'b \<Rightarrow> 'c \<Rightarrow> bool" where
   1.602 -  pred_compI [intro]: "r a b \<Longrightarrow> s b c \<Longrightarrow> (r OO s) a c"
   1.603 -
   1.604 -inductive_cases pred_compE [elim!]: "(r OO s) a c"
   1.605 -
   1.606 -lemma pred_comp_rel_comp_eq [pred_set_conv]:
   1.607 -  "((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)"
   1.608 -  by (auto simp add: fun_eq_iff)
   1.609 -
   1.610 -
   1.611 -subsubsection {* Converse *}
   1.612 -
   1.613 -inductive conversep :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> bool" ("(_^--1)" [1000] 1000)
   1.614 -  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.615 -  conversepI: "r a b \<Longrightarrow> r^--1 b a"
   1.616 -
   1.617 -notation (xsymbols)
   1.618 -  conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
   1.619 -
   1.620 -lemma conversepD:
   1.621 -  assumes ab: "r^--1 a b"
   1.622 -  shows "r b a" using ab
   1.623 -  by cases simp
   1.624 -
   1.625 -lemma conversep_iff [iff]: "r^--1 a b = r b a"
   1.626 -  by (iprover intro: conversepI dest: conversepD)
   1.627 -
   1.628 -lemma conversep_converse_eq [pred_set_conv]:
   1.629 -  "(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)"
   1.630 -  by (auto simp add: fun_eq_iff)
   1.631 -
   1.632 -lemma conversep_conversep [simp]: "(r^--1)^--1 = r"
   1.633 -  by (iprover intro: order_antisym conversepI dest: conversepD)
   1.634 -
   1.635 -lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
   1.636 -  by (iprover intro: order_antisym conversepI pred_compI
   1.637 -    elim: pred_compE dest: conversepD)
   1.638 -
   1.639 -lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1"
   1.640 -  by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
   1.641 -
   1.642 -lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1"
   1.643 -  by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
   1.644 -
   1.645 -lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>"
   1.646 -  by (auto simp add: fun_eq_iff)
   1.647 -
   1.648 -lemma conversep_eq [simp]: "(op =)^--1 = op ="
   1.649 -  by (auto simp add: fun_eq_iff)
   1.650 -
   1.651 -
   1.652 -subsubsection {* Domain *}
   1.653 -
   1.654 -inductive DomainP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
   1.655 -  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.656 -  DomainPI [intro]: "r a b \<Longrightarrow> DomainP r a"
   1.657 -
   1.658 -inductive_cases DomainPE [elim!]: "DomainP r a"
   1.659 -
   1.660 -lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
   1.661 -  by (blast intro!: Orderings.order_antisym predicate1I)
   1.662 -
   1.663 -
   1.664 -subsubsection {* Range *}
   1.665 -
   1.666 -inductive RangeP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool"
   1.667 -  for r :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where
   1.668 -  RangePI [intro]: "r a b \<Longrightarrow> RangeP r b"
   1.669 -
   1.670 -inductive_cases RangePE [elim!]: "RangeP r b"
   1.671 -
   1.672 -lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
   1.673 -  by (blast intro!: Orderings.order_antisym predicate1I)
   1.674 -
   1.675 -
   1.676 -subsubsection {* Inverse image *}
   1.677 -
   1.678 -definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
   1.679 -  "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
   1.680 -
   1.681 -lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
   1.682 -  by (simp add: inv_image_def inv_imagep_def)
   1.683 -
   1.684  lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
   1.685    by (simp add: inv_imagep_def)
   1.686  
   1.687 @@ -858,55 +892,5 @@
   1.688  
   1.689  lemmas Powp_mono [mono] = Pow_mono [to_pred]
   1.690  
   1.691 -
   1.692 -subsubsection {* Properties of predicate relations *}
   1.693 -
   1.694 -abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.695 -  "antisymP r \<equiv> antisym {(x, y). r x y}"
   1.696 -
   1.697 -abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.698 -  "transP r \<equiv> trans {(x, y). r x y}"
   1.699 -
   1.700 -abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   1.701 -  "single_valuedP r \<equiv> single_valued {(x, y). r x y}"
   1.702 -
   1.703 -(*FIXME inconsistencies: abbreviations vs. definitions, suffix `P` vs. suffix `p`*)
   1.704 -
   1.705 -definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.706 -  "reflp r \<longleftrightarrow> refl {(x, y). r x y}"
   1.707 -
   1.708 -definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.709 -  "symp r \<longleftrightarrow> sym {(x, y). r x y}"
   1.710 -
   1.711 -definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   1.712 -  "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   1.713 -
   1.714 -lemma reflpI:
   1.715 -  "(\<And>x. r x x) \<Longrightarrow> reflp r"
   1.716 -  by (auto intro: refl_onI simp add: reflp_def)
   1.717 -
   1.718 -lemma reflpE:
   1.719 -  assumes "reflp r"
   1.720 -  obtains "r x x"
   1.721 -  using assms by (auto dest: refl_onD simp add: reflp_def)
   1.722 -
   1.723 -lemma sympI:
   1.724 -  "(\<And>x y. r x y \<Longrightarrow> r y x) \<Longrightarrow> symp r"
   1.725 -  by (auto intro: symI simp add: symp_def)
   1.726 -
   1.727 -lemma sympE:
   1.728 -  assumes "symp r" and "r x y"
   1.729 -  obtains "r y x"
   1.730 -  using assms by (auto dest: symD simp add: symp_def)
   1.731 -
   1.732 -lemma transpI:
   1.733 -  "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   1.734 -  by (auto intro: transI simp add: transp_def)
   1.735 -  
   1.736 -lemma transpE:
   1.737 -  assumes "transp r" and "r x y" and "r y z"
   1.738 -  obtains "r x z"
   1.739 -  using assms by (auto dest: transD simp add: transp_def)
   1.740 -
   1.741  end
   1.742