| author | ballarin | 
| Fri, 14 Mar 2003 18:00:16 +0100 | |
| changeset 13864 | f44f121dd275 | 
| parent 13726 | 9550a6f4ed4a | 
| child 13867 | 1fdecd15437f | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Transitive_Closure.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 12691 | 7 | header {* Reflexive and Transitive closure of a relation *}
 | 
| 8 | ||
| 9 | theory Transitive_Closure = Inductive: | |
| 10 | ||
| 11 | text {*
 | |
| 12 |   @{text rtrancl} is reflexive/transitive closure,
 | |
| 13 |   @{text trancl} is transitive closure,
 | |
| 14 |   @{text reflcl} is reflexive closure.
 | |
| 15 | ||
| 16 |   These postfix operators have \emph{maximum priority}, forcing their
 | |
| 17 | operands to be atomic. | |
| 18 | *} | |
| 10213 | 19 | |
| 11327 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 20 | consts | 
| 12691 | 21 |   rtrancl :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_^*)" [1000] 999)
 | 
| 11327 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 22 | |
| 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 23 | inductive "r^*" | 
| 12691 | 24 | intros | 
| 12823 | 25 | rtrancl_refl [intro!, CPure.intro!, simp]: "(a, a) : r^*" | 
| 26 | rtrancl_into_rtrancl [CPure.intro]: "(a, b) : r^* ==> (b, c) : r ==> (a, c) : r^*" | |
| 11327 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 27 | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 28 | consts | 
| 12691 | 29 |   trancl :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_^+)" [1000] 999)
 | 
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 30 | |
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 31 | inductive "r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 32 | intros | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 33 | r_into_trancl [intro, CPure.intro]: "(a, b) : r ==> (a, b) : r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 34 | trancl_into_trancl [CPure.intro]: "(a, b) : r^+ ==> (b, c) : r ==> (a,c) : r^+" | 
| 10213 | 35 | |
| 36 | syntax | |
| 12691 | 37 |   "_reflcl" :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_^=)" [1000] 999)
 | 
| 10213 | 38 | translations | 
| 12691 | 39 | "r^=" == "r \<union> Id" | 
| 10213 | 40 | |
| 10827 | 41 | syntax (xsymbols) | 
| 12691 | 42 |   rtrancl :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_\\<^sup>*)" [1000] 999)
 | 
| 43 |   trancl :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_\\<^sup>+)" [1000] 999)
 | |
| 44 |   "_reflcl" :: "('a \<times> 'a) set => ('a \<times> 'a) set"    ("(_\\<^sup>=)" [1000] 999)
 | |
| 45 | ||
| 46 | ||
| 47 | subsection {* Reflexive-transitive closure *}
 | |
| 48 | ||
| 49 | lemma r_into_rtrancl [intro]: "!!p. p \<in> r ==> p \<in> r^*" | |
| 50 |   -- {* @{text rtrancl} of @{text r} contains @{text r} *}
 | |
| 51 | apply (simp only: split_tupled_all) | |
| 52 | apply (erule rtrancl_refl [THEN rtrancl_into_rtrancl]) | |
| 53 | done | |
| 54 | ||
| 55 | lemma rtrancl_mono: "r \<subseteq> s ==> r^* \<subseteq> s^*" | |
| 56 |   -- {* monotonicity of @{text rtrancl} *}
 | |
| 57 | apply (rule subsetI) | |
| 58 | apply (simp only: split_tupled_all) | |
| 59 | apply (erule rtrancl.induct) | |
| 60 | apply (rule_tac [2] rtrancl_into_rtrancl) | |
| 61 | apply blast+ | |
| 62 | done | |
| 63 | ||
| 12823 | 64 | theorem rtrancl_induct [consumes 1, induct set: rtrancl]: | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 65 | assumes a: "(a, b) : r^*" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 66 | and cases: "P a" "!!y z. [| (a, y) : r^*; (y, z) : r; P y |] ==> P z" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 67 | shows "P b" | 
| 12691 | 68 | proof - | 
| 69 | from a have "a = a --> P b" | |
| 12823 | 70 | by (induct "%x y. x = a --> P y" a b) (rules intro: cases)+ | 
| 12691 | 71 | thus ?thesis by rules | 
| 72 | qed | |
| 73 | ||
| 74 | ML_setup {*
 | |
| 75 |   bind_thm ("rtrancl_induct2", split_rule
 | |
| 76 |     (read_instantiate [("a","(ax,ay)"), ("b","(bx,by)")] (thm "rtrancl_induct")));
 | |
| 77 | *} | |
| 78 | ||
| 79 | lemma trans_rtrancl: "trans(r^*)" | |
| 80 |   -- {* transitivity of transitive closure!! -- by induction *}
 | |
| 12823 | 81 | proof (rule transI) | 
| 82 | fix x y z | |
| 83 | assume "(x, y) \<in> r\<^sup>*" | |
| 84 | assume "(y, z) \<in> r\<^sup>*" | |
| 85 | thus "(x, z) \<in> r\<^sup>*" by induct (rules!)+ | |
| 86 | qed | |
| 12691 | 87 | |
| 88 | lemmas rtrancl_trans = trans_rtrancl [THEN transD, standard] | |
| 89 | ||
| 90 | lemma rtranclE: | |
| 91 | "[| (a::'a,b) : r^*; (a = b) ==> P; | |
| 92 | !!y.[| (a,y) : r^*; (y,b) : r |] ==> P | |
| 93 | |] ==> P" | |
| 94 |   -- {* elimination of @{text rtrancl} -- by induction on a special formula *}
 | |
| 95 | proof - | |
| 96 | assume major: "(a::'a,b) : r^*" | |
| 97 | case rule_context | |
| 98 | show ?thesis | |
| 99 | apply (subgoal_tac "(a::'a) = b | (EX y. (a,y) : r^* & (y,b) : r)") | |
| 100 | apply (rule_tac [2] major [THEN rtrancl_induct]) | |
| 101 | prefer 2 apply (blast!) | |
| 102 | prefer 2 apply (blast!) | |
| 103 | apply (erule asm_rl exE disjE conjE prems)+ | |
| 104 | done | |
| 105 | qed | |
| 106 | ||
| 12823 | 107 | lemma converse_rtrancl_into_rtrancl: | 
| 108 | "(a, b) \<in> r \<Longrightarrow> (b, c) \<in> r\<^sup>* \<Longrightarrow> (a, c) \<in> r\<^sup>*" | |
| 109 | by (rule rtrancl_trans) rules+ | |
| 12691 | 110 | |
| 111 | text {*
 | |
| 112 |   \medskip More @{term "r^*"} equations and inclusions.
 | |
| 113 | *} | |
| 114 | ||
| 115 | lemma rtrancl_idemp [simp]: "(r^*)^* = r^*" | |
| 116 | apply auto | |
| 117 | apply (erule rtrancl_induct) | |
| 118 | apply (rule rtrancl_refl) | |
| 119 | apply (blast intro: rtrancl_trans) | |
| 120 | done | |
| 121 | ||
| 122 | lemma rtrancl_idemp_self_comp [simp]: "R^* O R^* = R^*" | |
| 123 | apply (rule set_ext) | |
| 124 | apply (simp only: split_tupled_all) | |
| 125 | apply (blast intro: rtrancl_trans) | |
| 126 | done | |
| 127 | ||
| 128 | lemma rtrancl_subset_rtrancl: "r \<subseteq> s^* ==> r^* \<subseteq> s^*" | |
| 129 | apply (drule rtrancl_mono) | |
| 130 | apply simp | |
| 131 | done | |
| 132 | ||
| 133 | lemma rtrancl_subset: "R \<subseteq> S ==> S \<subseteq> R^* ==> S^* = R^*" | |
| 134 | apply (drule rtrancl_mono) | |
| 135 | apply (drule rtrancl_mono) | |
| 136 | apply simp | |
| 137 | apply blast | |
| 138 | done | |
| 139 | ||
| 140 | lemma rtrancl_Un_rtrancl: "(R^* \<union> S^*)^* = (R \<union> S)^*" | |
| 141 | by (blast intro!: rtrancl_subset intro: r_into_rtrancl rtrancl_mono [THEN subsetD]) | |
| 142 | ||
| 143 | lemma rtrancl_reflcl [simp]: "(R^=)^* = R^*" | |
| 144 | by (blast intro!: rtrancl_subset intro: r_into_rtrancl) | |
| 145 | ||
| 146 | lemma rtrancl_r_diff_Id: "(r - Id)^* = r^*" | |
| 147 | apply (rule sym) | |
| 148 | apply (rule rtrancl_subset) | |
| 149 | apply blast | |
| 150 | apply clarify | |
| 151 | apply (rename_tac a b) | |
| 152 | apply (case_tac "a = b") | |
| 153 | apply blast | |
| 154 | apply (blast intro!: r_into_rtrancl) | |
| 155 | done | |
| 156 | ||
| 12823 | 157 | theorem rtrancl_converseD: | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 158 | assumes r: "(x, y) \<in> (r^-1)^*" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 159 | shows "(y, x) \<in> r^*" | 
| 12823 | 160 | proof - | 
| 161 | from r show ?thesis | |
| 162 | by induct (rules intro: rtrancl_trans dest!: converseD)+ | |
| 163 | qed | |
| 12691 | 164 | |
| 12823 | 165 | theorem rtrancl_converseI: | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 166 | assumes r: "(y, x) \<in> r^*" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 167 | shows "(x, y) \<in> (r^-1)^*" | 
| 12823 | 168 | proof - | 
| 169 | from r show ?thesis | |
| 170 | by induct (rules intro: rtrancl_trans converseI)+ | |
| 171 | qed | |
| 12691 | 172 | |
| 173 | lemma rtrancl_converse: "(r^-1)^* = (r^*)^-1" | |
| 174 | by (fast dest!: rtrancl_converseD intro!: rtrancl_converseI) | |
| 175 | ||
| 12823 | 176 | theorem converse_rtrancl_induct: | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 177 | assumes major: "(a, b) : r^*" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 178 | and cases: "P b" "!!y z. [| (y, z) : r; (z, b) : r^*; P z |] ==> P y" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12823diff
changeset | 179 | shows "P a" | 
| 12691 | 180 | proof - | 
| 12823 | 181 | from rtrancl_converseI [OF major] | 
| 12691 | 182 | show ?thesis | 
| 12823 | 183 | by induct (rules intro: cases dest!: converseD rtrancl_converseD)+ | 
| 12691 | 184 | qed | 
| 185 | ||
| 186 | ML_setup {*
 | |
| 187 |   bind_thm ("converse_rtrancl_induct2", split_rule
 | |
| 188 |     (read_instantiate [("a","(ax,ay)"),("b","(bx,by)")] (thm "converse_rtrancl_induct")));
 | |
| 189 | *} | |
| 190 | ||
| 191 | lemma converse_rtranclE: | |
| 192 | "[| (x,z):r^*; | |
| 193 | x=z ==> P; | |
| 194 | !!y. [| (x,y):r; (y,z):r^* |] ==> P | |
| 195 | |] ==> P" | |
| 196 | proof - | |
| 197 | assume major: "(x,z):r^*" | |
| 198 | case rule_context | |
| 199 | show ?thesis | |
| 200 | apply (subgoal_tac "x = z | (EX y. (x,y) : r & (y,z) : r^*)") | |
| 201 | apply (rule_tac [2] major [THEN converse_rtrancl_induct]) | |
| 13726 | 202 | prefer 2 apply rules | 
| 203 | prefer 2 apply rules | |
| 12691 | 204 | apply (erule asm_rl exE disjE conjE prems)+ | 
| 205 | done | |
| 206 | qed | |
| 207 | ||
| 208 | ML_setup {*
 | |
| 209 |   bind_thm ("converse_rtranclE2", split_rule
 | |
| 210 |     (read_instantiate [("x","(xa,xb)"), ("z","(za,zb)")] (thm "converse_rtranclE")));
 | |
| 211 | *} | |
| 212 | ||
| 213 | lemma r_comp_rtrancl_eq: "r O r^* = r^* O r" | |
| 214 | by (blast elim: rtranclE converse_rtranclE | |
| 215 | intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl) | |
| 216 | ||
| 217 | ||
| 218 | subsection {* Transitive closure *}
 | |
| 10331 | 219 | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 220 | lemma trancl_mono: "!!p. p \<in> r^+ ==> r \<subseteq> s ==> p \<in> s^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 221 | apply (simp only: split_tupled_all) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 222 | apply (erule trancl.induct) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 223 | apply (rules dest: subsetD)+ | 
| 12691 | 224 | done | 
| 225 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 226 | lemma r_into_trancl': "!!p. p : r ==> p : r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 227 | by (simp only: split_tupled_all) (erule r_into_trancl) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 228 | |
| 12691 | 229 | text {*
 | 
| 230 |   \medskip Conversions between @{text trancl} and @{text rtrancl}.
 | |
| 231 | *} | |
| 232 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 233 | lemma trancl_into_rtrancl: "(a, b) \<in> r^+ ==> (a, b) \<in> r^*" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 234 | by (erule trancl.induct) rules+ | 
| 12691 | 235 | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 236 | lemma rtrancl_into_trancl1: assumes r: "(a, b) \<in> r^*" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 237 | shows "!!c. (b, c) \<in> r ==> (a, c) \<in> r^+" using r | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 238 | by induct rules+ | 
| 12691 | 239 | |
| 240 | lemma rtrancl_into_trancl2: "[| (a,b) : r; (b,c) : r^* |] ==> (a,c) : r^+" | |
| 241 |   -- {* intro rule from @{text r} and @{text rtrancl} *}
 | |
| 242 | apply (erule rtranclE) | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 243 | apply rules | 
| 12691 | 244 | apply (rule rtrancl_trans [THEN rtrancl_into_trancl1]) | 
| 245 | apply (assumption | rule r_into_rtrancl)+ | |
| 246 | done | |
| 247 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 248 | lemma trancl_induct [consumes 1, induct set: trancl]: | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 249 | assumes a: "(a,b) : r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 250 | and cases: "!!y. (a, y) : r ==> P y" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 251 | "!!y z. (a,y) : r^+ ==> (y, z) : r ==> P y ==> P z" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 252 | shows "P b" | 
| 12691 | 253 |   -- {* Nice induction rule for @{text trancl} *}
 | 
| 254 | proof - | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 255 | from a have "a = a --> P b" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 256 | by (induct "%x y. x = a --> P y" a b) (rules intro: cases)+ | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 257 | thus ?thesis by rules | 
| 12691 | 258 | qed | 
| 259 | ||
| 260 | lemma trancl_trans_induct: | |
| 261 | "[| (x,y) : r^+; | |
| 262 | !!x y. (x,y) : r ==> P x y; | |
| 263 | !!x y z. [| (x,y) : r^+; P x y; (y,z) : r^+; P y z |] ==> P x z | |
| 264 | |] ==> P x y" | |
| 265 |   -- {* Another induction rule for trancl, incorporating transitivity *}
 | |
| 266 | proof - | |
| 267 | assume major: "(x,y) : r^+" | |
| 268 | case rule_context | |
| 269 | show ?thesis | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 270 | by (rules intro: r_into_trancl major [THEN trancl_induct] prems) | 
| 12691 | 271 | qed | 
| 272 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 273 | inductive_cases tranclE: "(a, b) : r^+" | 
| 10980 | 274 | |
| 12691 | 275 | lemma trans_trancl: "trans(r^+)" | 
| 276 |   -- {* Transitivity of @{term "r^+"} *}
 | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 277 | proof (rule transI) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 278 | fix x y z | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 279 | assume "(x, y) \<in> r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 280 | assume "(y, z) \<in> r^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 281 | thus "(x, z) \<in> r^+" by induct (rules!)+ | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 282 | qed | 
| 12691 | 283 | |
| 284 | lemmas trancl_trans = trans_trancl [THEN transD, standard] | |
| 285 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 286 | lemma rtrancl_trancl_trancl: assumes r: "(x, y) \<in> r^*" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 287 | shows "!!z. (y, z) \<in> r^+ ==> (x, z) \<in> r^+" using r | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 288 | by induct (rules intro: trancl_trans)+ | 
| 12691 | 289 | |
| 290 | lemma trancl_into_trancl2: "(a, b) \<in> r ==> (b, c) \<in> r^+ ==> (a, c) \<in> r^+" | |
| 291 | by (erule transD [OF trans_trancl r_into_trancl]) | |
| 292 | ||
| 293 | lemma trancl_insert: | |
| 294 |   "(insert (y, x) r)^+ = r^+ \<union> {(a, b). (a, y) \<in> r^* \<and> (x, b) \<in> r^*}"
 | |
| 295 |   -- {* primitive recursion for @{text trancl} over finite relations *}
 | |
| 296 | apply (rule equalityI) | |
| 297 | apply (rule subsetI) | |
| 298 | apply (simp only: split_tupled_all) | |
| 299 | apply (erule trancl_induct) | |
| 300 | apply blast | |
| 301 | apply (blast intro: rtrancl_into_trancl1 trancl_into_rtrancl r_into_trancl trancl_trans) | |
| 302 | apply (rule subsetI) | |
| 303 | apply (blast intro: trancl_mono rtrancl_mono | |
| 304 | [THEN [2] rev_subsetD] rtrancl_trancl_trancl rtrancl_into_trancl2) | |
| 305 | done | |
| 306 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 307 | lemma trancl_converseI: "(x, y) \<in> (r^+)^-1 ==> (x, y) \<in> (r^-1)^+" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 308 | apply (drule converseD) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 309 | apply (erule trancl.induct) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 310 | apply (rules intro: converseI trancl_trans)+ | 
| 12691 | 311 | done | 
| 312 | ||
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 313 | lemma trancl_converseD: "(x, y) \<in> (r^-1)^+ ==> (x, y) \<in> (r^+)^-1" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 314 | apply (rule converseI) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 315 | apply (erule trancl.induct) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 316 | apply (rules dest: converseD intro: trancl_trans)+ | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 317 | done | 
| 12691 | 318 | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 319 | lemma trancl_converse: "(r^-1)^+ = (r^+)^-1" | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 320 | by (fastsimp simp add: split_tupled_all | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 321 | intro!: trancl_converseI trancl_converseD) | 
| 12691 | 322 | |
| 323 | lemma converse_trancl_induct: | |
| 324 | "[| (a,b) : r^+; !!y. (y,b) : r ==> P(y); | |
| 325 | !!y z.[| (y,z) : r; (z,b) : r^+; P(z) |] ==> P(y) |] | |
| 326 | ==> P(a)" | |
| 327 | proof - | |
| 328 | assume major: "(a,b) : r^+" | |
| 329 | case rule_context | |
| 330 | show ?thesis | |
| 331 | apply (rule major [THEN converseI, THEN trancl_converseI [THEN trancl_induct]]) | |
| 332 | apply (rule prems) | |
| 333 | apply (erule converseD) | |
| 334 | apply (blast intro: prems dest!: trancl_converseD) | |
| 335 | done | |
| 336 | qed | |
| 337 | ||
| 338 | lemma tranclD: "(x, y) \<in> R^+ ==> EX z. (x, z) \<in> R \<and> (z, y) \<in> R^*" | |
| 339 | apply (erule converse_trancl_induct) | |
| 340 | apply auto | |
| 341 | apply (blast intro: rtrancl_trans) | |
| 342 | done | |
| 343 | ||
| 344 | lemma irrefl_tranclI: "r^-1 \<inter> r^+ = {} ==> (x, x) \<notin> r^+"
 | |
| 345 | apply (subgoal_tac "ALL y. (x, y) : r^+ --> x \<noteq> y") | |
| 346 | apply fast | |
| 347 | apply (intro strip) | |
| 348 | apply (erule trancl_induct) | |
| 349 | apply (auto intro: r_into_trancl) | |
| 350 | done | |
| 351 | ||
| 352 | lemma irrefl_trancl_rD: "!!X. ALL x. (x, x) \<notin> r^+ ==> (x, y) \<in> r ==> x \<noteq> y" | |
| 353 | by (blast dest: r_into_trancl) | |
| 354 | ||
| 355 | lemma trancl_subset_Sigma_aux: | |
| 356 | "(a, b) \<in> r^* ==> r \<subseteq> A \<times> A ==> a = b \<or> a \<in> A" | |
| 357 | apply (erule rtrancl_induct) | |
| 358 | apply auto | |
| 359 | done | |
| 360 | ||
| 361 | lemma trancl_subset_Sigma: "r \<subseteq> A \<times> A ==> r^+ \<subseteq> A \<times> A" | |
| 13704 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 362 | apply (rule subsetI) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 363 | apply (simp only: split_tupled_all) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 364 | apply (erule tranclE) | 
| 
854501b1e957
Transitive closure is now defined inductively as well.
 berghofe parents: 
12937diff
changeset | 365 | apply (blast dest!: trancl_into_rtrancl trancl_subset_Sigma_aux)+ | 
| 12691 | 366 | done | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 367 | |
| 11090 | 368 | lemma reflcl_trancl [simp]: "(r^+)^= = r^*" | 
| 11084 | 369 | apply safe | 
| 12691 | 370 | apply (erule trancl_into_rtrancl) | 
| 11084 | 371 | apply (blast elim: rtranclE dest: rtrancl_into_trancl1) | 
| 372 | done | |
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 373 | |
| 11090 | 374 | lemma trancl_reflcl [simp]: "(r^=)^+ = r^*" | 
| 11084 | 375 | apply safe | 
| 376 | apply (drule trancl_into_rtrancl) | |
| 377 | apply simp | |
| 378 | apply (erule rtranclE) | |
| 379 | apply safe | |
| 380 | apply (rule r_into_trancl) | |
| 381 | apply simp | |
| 382 | apply (rule rtrancl_into_trancl1) | |
| 383 | apply (erule rtrancl_reflcl [THEN equalityD2, THEN subsetD]) | |
| 384 | apply fast | |
| 385 | done | |
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 386 | |
| 11090 | 387 | lemma trancl_empty [simp]: "{}^+ = {}"
 | 
| 11084 | 388 | by (auto elim: trancl_induct) | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 389 | |
| 11090 | 390 | lemma rtrancl_empty [simp]: "{}^* = Id"
 | 
| 11084 | 391 | by (rule subst [OF reflcl_trancl]) simp | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 392 | |
| 11090 | 393 | lemma rtranclD: "(a, b) \<in> R^* ==> a = b \<or> a \<noteq> b \<and> (a, b) \<in> R^+" | 
| 11084 | 394 | by (force simp add: reflcl_trancl [symmetric] simp del: reflcl_trancl) | 
| 395 | ||
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 396 | |
| 12691 | 397 | text {* @{text Domain} and @{text Range} *}
 | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 398 | |
| 11090 | 399 | lemma Domain_rtrancl [simp]: "Domain (R^*) = UNIV" | 
| 11084 | 400 | by blast | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 401 | |
| 11090 | 402 | lemma Range_rtrancl [simp]: "Range (R^*) = UNIV" | 
| 11084 | 403 | by blast | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 404 | |
| 11090 | 405 | lemma rtrancl_Un_subset: "(R^* \<union> S^*) \<subseteq> (R Un S)^*" | 
| 11084 | 406 | by (rule rtrancl_Un_rtrancl [THEN subst]) fast | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 407 | |
| 11090 | 408 | lemma in_rtrancl_UnI: "x \<in> R^* \<or> x \<in> S^* ==> x \<in> (R \<union> S)^*" | 
| 11084 | 409 | by (blast intro: subsetD [OF rtrancl_Un_subset]) | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 410 | |
| 11090 | 411 | lemma trancl_domain [simp]: "Domain (r^+) = Domain r" | 
| 11084 | 412 | by (unfold Domain_def) (blast dest: tranclD) | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 413 | |
| 11090 | 414 | lemma trancl_range [simp]: "Range (r^+) = Range r" | 
| 11084 | 415 | by (simp add: Range_def trancl_converse [symmetric]) | 
| 10996 
74e970389def
Moved some thms from Transitive_ClosureTr.ML to Transitive_Closure.thy
 nipkow parents: 
10980diff
changeset | 416 | |
| 11115 | 417 | lemma Not_Domain_rtrancl: | 
| 12691 | 418 | "x ~: Domain R ==> ((x, y) : R^*) = (x = y)" | 
| 419 | apply auto | |
| 420 | by (erule rev_mp, erule rtrancl_induct, auto) | |
| 421 | ||
| 11327 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 422 | |
| 12691 | 423 | text {* More about converse @{text rtrancl} and @{text trancl}, should
 | 
| 424 | be merged with main body. *} | |
| 12428 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 425 | |
| 12691 | 426 | lemma r_r_into_trancl: "(a, b) \<in> R ==> (b, c) \<in> R ==> (a, c) \<in> R^+" | 
| 12428 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 427 | by (fast intro: trancl_trans) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 428 | |
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 429 | lemma trancl_into_trancl [rule_format]: | 
| 12691 | 430 | "(a, b) \<in> r\<^sup>+ ==> (b, c) \<in> r --> (a,c) \<in> r\<^sup>+" | 
| 431 | apply (erule trancl_induct) | |
| 12428 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 432 | apply (fast intro: r_r_into_trancl) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 433 | apply (fast intro: r_r_into_trancl trancl_trans) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 434 | done | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 435 | |
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 436 | lemma trancl_rtrancl_trancl: | 
| 12691 | 437 | "(a, b) \<in> r\<^sup>+ ==> (b, c) \<in> r\<^sup>* ==> (a, c) \<in> r\<^sup>+" | 
| 12428 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 438 | apply (drule tranclD) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 439 | apply (erule exE, erule conjE) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 440 | apply (drule rtrancl_trans, assumption) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 441 | apply (drule rtrancl_into_trancl2, assumption) | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 442 | apply assumption | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 443 | done | 
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 444 | |
| 12691 | 445 | lemmas transitive_closure_trans [trans] = | 
| 446 | r_r_into_trancl trancl_trans rtrancl_trans | |
| 447 | trancl_into_trancl trancl_into_trancl2 | |
| 448 | rtrancl_into_rtrancl converse_rtrancl_into_rtrancl | |
| 449 | rtrancl_trancl_trancl trancl_rtrancl_trancl | |
| 12428 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 450 | |
| 
f3033eed309a
setup [trans] rules for calculational Isar reasoning
 kleing parents: 
11327diff
changeset | 451 | declare trancl_into_rtrancl [elim] | 
| 11327 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 452 | |
| 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 453 | declare rtranclE [cases set: rtrancl] | 
| 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 454 | declare tranclE [cases set: trancl] | 
| 
cd2c27a23df1
Transitive closure is now defined via "inductive".
 berghofe parents: 
11115diff
changeset | 455 | |
| 10213 | 456 | end |