| author | paulson | 
| Thu, 15 May 1997 12:53:12 +0200 | |
| changeset 3194 | 36bfceef1800 | 
| parent 2935 | 998cb95fdd43 | 
| child 3413 | c1f63cc3a768 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/trancl | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 6 | For trancl.thy. Theorems about the transitive closure of a relation | |
| 7 | *) | |
| 8 | ||
| 9 | open Trancl; | |
| 10 | ||
| 11 | (** The relation rtrancl **) | |
| 12 | ||
| 13 | goal Trancl.thy "mono(%s. id Un (r O s))"; | |
| 14 | by (rtac monoI 1); | |
| 15 | by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1)); | |
| 16 | qed "rtrancl_fun_mono"; | |
| 17 | ||
| 18 | val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); | |
| 19 | ||
| 20 | (*Reflexivity of rtrancl*) | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 21 | goal Trancl.thy "(a,a) : r^*"; | 
| 923 | 22 | by (stac rtrancl_unfold 1); | 
| 2891 | 23 | by (Blast_tac 1); | 
| 923 | 24 | qed "rtrancl_refl"; | 
| 25 | ||
| 1921 | 26 | Addsimps [rtrancl_refl]; | 
| 27 | AddSIs [rtrancl_refl]; | |
| 28 | ||
| 29 | ||
| 923 | 30 | (*Closure under composition with r*) | 
| 1921 | 31 | goal Trancl.thy "!!r. [| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^*"; | 
| 923 | 32 | by (stac rtrancl_unfold 1); | 
| 2891 | 33 | by (Blast_tac 1); | 
| 923 | 34 | qed "rtrancl_into_rtrancl"; | 
| 35 | ||
| 36 | (*rtrancl of r contains r*) | |
| 1301 | 37 | goal Trancl.thy "!!p. p : r ==> p : r^*"; | 
| 1552 | 38 | by (split_all_tac 1); | 
| 1301 | 39 | by (etac (rtrancl_refl RS rtrancl_into_rtrancl) 1); | 
| 923 | 40 | qed "r_into_rtrancl"; | 
| 41 | ||
| 42 | (*monotonicity of rtrancl*) | |
| 43 | goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*"; | |
| 1552 | 44 | by (REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1)); | 
| 923 | 45 | qed "rtrancl_mono"; | 
| 46 | ||
| 47 | (** standard induction rule **) | |
| 48 | ||
| 49 | val major::prems = goal Trancl.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 50 | "[| (a,b) : r^*; \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 51 | \ !!x. P((x,x)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 52 | \ !!x y z.[| P((x,y)); (x,y): r^*; (y,z): r |] ==> P((x,z)) |] \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 53 | \ ==> P((a,b))"; | 
| 923 | 54 | by (rtac ([rtrancl_def, rtrancl_fun_mono, major] MRS def_induct) 1); | 
| 2935 | 55 | by (blast_tac (!claset addIs prems) 1); | 
| 923 | 56 | qed "rtrancl_full_induct"; | 
| 57 | ||
| 58 | (*nice induction rule*) | |
| 59 | val major::prems = goal Trancl.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 60 | "[| (a::'a,b) : r^*; \ | 
| 923 | 61 | \ P(a); \ | 
| 1465 | 62 | \ !!y z.[| (a,y) : r^*; (y,z) : r; P(y) |] ==> P(z) |] \ | 
| 923 | 63 | \ ==> P(b)"; | 
| 64 | (*by induction on this formula*) | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 65 | by (subgoal_tac "! y. (a::'a,b) = (a,y) --> P(y)" 1); | 
| 923 | 66 | (*now solve first subgoal: this formula is sufficient*) | 
| 2891 | 67 | by (Blast_tac 1); | 
| 923 | 68 | (*now do the induction*) | 
| 69 | by (resolve_tac [major RS rtrancl_full_induct] 1); | |
| 2935 | 70 | by (blast_tac (!claset addIs prems) 1); | 
| 71 | by (blast_tac (!claset addIs prems) 1); | |
| 923 | 72 | qed "rtrancl_induct"; | 
| 73 | ||
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1706diff
changeset | 74 | bind_thm | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1706diff
changeset | 75 |   ("rtrancl_induct2",
 | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1706diff
changeset | 76 | Prod_Syntax.split_rule | 
| 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1706diff
changeset | 77 |      (read_instantiate [("a","(ax,ay)"), ("b","(bx,by)")] rtrancl_induct));
 | 
| 1706 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 78 | |
| 923 | 79 | (*transitivity of transitive closure!! -- by induction.*) | 
| 1642 | 80 | goalw Trancl.thy [trans_def] "trans(r^*)"; | 
| 1786 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 berghofe parents: 
1766diff
changeset | 81 | by (safe_tac (!claset)); | 
| 1642 | 82 | by (eres_inst_tac [("b","z")] rtrancl_induct 1);
 | 
| 2922 | 83 | by (ALLGOALS(blast_tac (!claset addIs [rtrancl_into_rtrancl]))); | 
| 1642 | 84 | qed "trans_rtrancl"; | 
| 85 | ||
| 86 | bind_thm ("rtrancl_trans", trans_rtrancl RS transD);
 | |
| 87 | ||
| 923 | 88 | |
| 89 | (*elimination of rtrancl -- by induction on a special formula*) | |
| 90 | val major::prems = goal Trancl.thy | |
| 1465 | 91 | "[| (a::'a,b) : r^*; (a = b) ==> P; \ | 
| 92 | \ !!y.[| (a,y) : r^*; (y,b) : r |] ==> P \ | |
| 923 | 93 | \ |] ==> P"; | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 94 | by (subgoal_tac "(a::'a) = b | (? y. (a,y) : r^* & (y,b) : r)" 1); | 
| 923 | 95 | by (rtac (major RS rtrancl_induct) 2); | 
| 2935 | 96 | by (blast_tac (!claset addIs prems) 2); | 
| 97 | by (blast_tac (!claset addIs prems) 2); | |
| 923 | 98 | by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); | 
| 99 | qed "rtranclE"; | |
| 100 | ||
| 1642 | 101 | bind_thm ("rtrancl_into_rtrancl2", r_into_rtrancl RS rtrancl_trans);
 | 
| 102 | ||
| 103 | ||
| 104 | (*** More r^* equations and inclusions ***) | |
| 105 | ||
| 106 | goal Trancl.thy "(r^*)^* = r^*"; | |
| 107 | by (rtac set_ext 1); | |
| 108 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 109 | by (hyp_subst_tac 1); | |
| 110 | by (rtac iffI 1); | |
| 1552 | 111 | by (etac rtrancl_induct 1); | 
| 1642 | 112 | by (rtac rtrancl_refl 1); | 
| 2922 | 113 | by (blast_tac (!claset addIs [rtrancl_trans]) 1); | 
| 1642 | 114 | by (etac r_into_rtrancl 1); | 
| 115 | qed "rtrancl_idemp"; | |
| 116 | Addsimps [rtrancl_idemp]; | |
| 117 | ||
| 118 | goal Trancl.thy "!!r s. r <= s^* ==> r^* <= s^*"; | |
| 2031 | 119 | by (dtac rtrancl_mono 1); | 
| 1642 | 120 | by (Asm_full_simp_tac 1); | 
| 121 | qed "rtrancl_subset_rtrancl"; | |
| 122 | ||
| 123 | goal Trancl.thy "!!R. [| R <= S; S <= R^* |] ==> S^* = R^*"; | |
| 124 | by (dtac rtrancl_mono 1); | |
| 125 | by (dtac rtrancl_mono 1); | |
| 126 | by (Asm_full_simp_tac 1); | |
| 2891 | 127 | by (Blast_tac 1); | 
| 1642 | 128 | qed "rtrancl_subset"; | 
| 129 | ||
| 130 | goal Trancl.thy "!!R. (R^* Un S^*)^* = (R Un S)^*"; | |
| 2922 | 131 | by (blast_tac (!claset addSIs [rtrancl_subset] | 
| 132 | addIs [r_into_rtrancl, rtrancl_mono RS subsetD]) 1); | |
| 1642 | 133 | qed "rtrancl_Un_rtrancl"; | 
| 1496 | 134 | |
| 1642 | 135 | goal Trancl.thy "(R^=)^* = R^*"; | 
| 2922 | 136 | by (blast_tac (!claset addSIs [rtrancl_subset] | 
| 137 | addIs [rtrancl_refl, r_into_rtrancl]) 1); | |
| 1642 | 138 | qed "rtrancl_reflcl"; | 
| 139 | Addsimps [rtrancl_reflcl]; | |
| 140 | ||
| 141 | goal Trancl.thy "!!r. (x,y) : (converse r)^* ==> (x,y) : converse(r^*)"; | |
| 142 | by (rtac converseI 1); | |
| 143 | by (etac rtrancl_induct 1); | |
| 144 | by (rtac rtrancl_refl 1); | |
| 2922 | 145 | by (blast_tac (!claset addIs [r_into_rtrancl,rtrancl_trans]) 1); | 
| 1642 | 146 | qed "rtrancl_converseD"; | 
| 147 | ||
| 148 | goal Trancl.thy "!!r. (x,y) : converse(r^*) ==> (x,y) : (converse r)^*"; | |
| 149 | by (dtac converseD 1); | |
| 150 | by (etac rtrancl_induct 1); | |
| 151 | by (rtac rtrancl_refl 1); | |
| 2922 | 152 | by (blast_tac (!claset addIs [r_into_rtrancl,rtrancl_trans]) 1); | 
| 1642 | 153 | qed "rtrancl_converseI"; | 
| 154 | ||
| 155 | goal Trancl.thy "(converse r)^* = converse(r^*)"; | |
| 1786 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 berghofe parents: 
1766diff
changeset | 156 | by (safe_tac (!claset addSIs [rtrancl_converseI])); | 
| 1642 | 157 | by (res_inst_tac [("p","x")] PairE 1);
 | 
| 158 | by (hyp_subst_tac 1); | |
| 159 | by (etac rtrancl_converseD 1); | |
| 160 | qed "rtrancl_converse"; | |
| 161 | ||
| 1706 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 162 | val major::prems = goal Trancl.thy | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 163 | "[| (a,b) : r^*; P(b); \ | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 164 | \ !!y z.[| (y,z) : r; (z,b) : r^*; P(z) |] ==> P(y) |] \ | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 165 | \ ==> P(a)"; | 
| 2031 | 166 | by (rtac ((major RS converseI RS rtrancl_converseI) RS rtrancl_induct) 1); | 
| 167 | by (resolve_tac prems 1); | |
| 2935 | 168 | by (blast_tac (!claset addIs prems addSDs[rtrancl_converseD])1); | 
| 1706 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 169 | qed "converse_rtrancl_induct"; | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 170 | |
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 171 | val prems = goal Trancl.thy | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 172 | "[| ((a,b),(c,d)) : r^*; P c d; \ | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 173 | \ !!x y z u.[| ((x,y),(z,u)) : r; ((z,u),(c,d)) : r^*; P z u |] ==> P x y\ | 
| 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 174 | \ |] ==> P a b"; | 
| 2031 | 175 | by (res_inst_tac[("R","P")]splitD 1);
 | 
| 176 | by (res_inst_tac[("P","split P")]converse_rtrancl_induct 1);
 | |
| 177 | by (resolve_tac prems 1); | |
| 178 | by (Simp_tac 1); | |
| 179 | by (resolve_tac prems 1); | |
| 180 | by (split_all_tac 1); | |
| 181 | by (Asm_full_simp_tac 1); | |
| 182 | by (REPEAT(ares_tac prems 1)); | |
| 1706 
4e0d5c7f57fa
Added backwards rtrancl_induct and special versions for pairs.
 nipkow parents: 
1642diff
changeset | 183 | qed "converse_rtrancl_induct2"; | 
| 1496 | 184 | |
| 923 | 185 | |
| 186 | (**** The relation trancl ****) | |
| 187 | ||
| 188 | (** Conversions between trancl and rtrancl **) | |
| 189 | ||
| 190 | val [major] = goalw Trancl.thy [trancl_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 191 | "(a,b) : r^+ ==> (a,b) : r^*"; | 
| 923 | 192 | by (resolve_tac [major RS compEpair] 1); | 
| 193 | by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); | |
| 194 | qed "trancl_into_rtrancl"; | |
| 195 | ||
| 196 | (*r^+ contains r*) | |
| 197 | val [prem] = goalw Trancl.thy [trancl_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 198 | "[| (a,b) : r |] ==> (a,b) : r^+"; | 
| 923 | 199 | by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); | 
| 200 | qed "r_into_trancl"; | |
| 201 | ||
| 202 | (*intro rule by definition: from rtrancl and r*) | |
| 203 | val prems = goalw Trancl.thy [trancl_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 204 | "[| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^+"; | 
| 923 | 205 | by (REPEAT (resolve_tac ([compI]@prems) 1)); | 
| 206 | qed "rtrancl_into_trancl1"; | |
| 207 | ||
| 208 | (*intro rule from r and rtrancl*) | |
| 209 | val prems = goal Trancl.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 210 | "[| (a,b) : r; (b,c) : r^* |] ==> (a,c) : r^+"; | 
| 923 | 211 | by (resolve_tac (prems RL [rtranclE]) 1); | 
| 212 | by (etac subst 1); | |
| 213 | by (resolve_tac (prems RL [r_into_trancl]) 1); | |
| 1122 
20b708827030
renamed trans_rtrancl to rtrancl_trans and modified it by expanding trans.
 nipkow parents: 
1121diff
changeset | 214 | by (rtac (rtrancl_trans RS rtrancl_into_trancl1) 1); | 
| 923 | 215 | by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1)); | 
| 216 | qed "rtrancl_into_trancl2"; | |
| 217 | ||
| 1642 | 218 | (*Nice induction rule for trancl*) | 
| 219 | val major::prems = goal Trancl.thy | |
| 220 | "[| (a,b) : r^+; \ | |
| 221 | \ !!y. [| (a,y) : r |] ==> P(y); \ | |
| 222 | \ !!y z.[| (a,y) : r^+; (y,z) : r; P(y) |] ==> P(z) \ | |
| 223 | \ |] ==> P(b)"; | |
| 224 | by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); | |
| 225 | (*by induction on this formula*) | |
| 226 | by (subgoal_tac "ALL z. (y,z) : r --> P(z)" 1); | |
| 227 | (*now solve first subgoal: this formula is sufficient*) | |
| 2891 | 228 | by (Blast_tac 1); | 
| 1642 | 229 | by (etac rtrancl_induct 1); | 
| 2935 | 230 | by (ALLGOALS (blast_tac (!claset addIs (rtrancl_into_trancl1::prems)))); | 
| 1642 | 231 | qed "trancl_induct"; | 
| 232 | ||
| 923 | 233 | (*elimination of r^+ -- NOT an induction rule*) | 
| 234 | val major::prems = goal Trancl.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 235 | "[| (a::'a,b) : r^+; \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 236 | \ (a,b) : r ==> P; \ | 
| 1465 | 237 | \ !!y.[| (a,y) : r^+; (y,b) : r |] ==> P \ | 
| 923 | 238 | \ |] ==> P"; | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 239 | by (subgoal_tac "(a::'a,b) : r | (? y. (a,y) : r^+ & (y,b) : r)" 1); | 
| 923 | 240 | by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1)); | 
| 241 | by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); | |
| 242 | by (etac rtranclE 1); | |
| 2891 | 243 | by (Blast_tac 1); | 
| 2922 | 244 | by (blast_tac (!claset addSIs [rtrancl_into_trancl1]) 1); | 
| 923 | 245 | qed "tranclE"; | 
| 246 | ||
| 247 | (*Transitivity of r^+. | |
| 248 | Proved by unfolding since it uses transitivity of rtrancl. *) | |
| 249 | goalw Trancl.thy [trancl_def] "trans(r^+)"; | |
| 250 | by (rtac transI 1); | |
| 251 | by (REPEAT (etac compEpair 1)); | |
| 1122 
20b708827030
renamed trans_rtrancl to rtrancl_trans and modified it by expanding trans.
 nipkow parents: 
1121diff
changeset | 252 | by (rtac (rtrancl_into_rtrancl RS (rtrancl_trans RS compI)) 1); | 
| 923 | 253 | by (REPEAT (assume_tac 1)); | 
| 254 | qed "trans_trancl"; | |
| 255 | ||
| 1642 | 256 | bind_thm ("trancl_trans", trans_trancl RS transD);
 | 
| 257 | ||
| 923 | 258 | val prems = goal Trancl.thy | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 259 | "[| (a,b) : r; (b,c) : r^+ |] ==> (a,c) : r^+"; | 
| 923 | 260 | by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1); | 
| 261 | by (resolve_tac prems 1); | |
| 262 | by (resolve_tac prems 1); | |
| 263 | qed "trancl_into_trancl2"; | |
| 264 | ||
| 1130 | 265 | |
| 923 | 266 | val major::prems = goal Trancl.thy | 
| 1642 | 267 | "[| (a,b) : r^*; r <= A Times A |] ==> a=b | a:A"; | 
| 923 | 268 | by (cut_facts_tac prems 1); | 
| 269 | by (rtac (major RS rtrancl_induct) 1); | |
| 270 | by (rtac (refl RS disjI1) 1); | |
| 2891 | 271 | by (Blast_tac 1); | 
| 1642 | 272 | val lemma = result(); | 
| 923 | 273 | |
| 274 | goalw Trancl.thy [trancl_def] | |
| 1642 | 275 | "!!r. r <= A Times A ==> r^+ <= A Times A"; | 
| 2891 | 276 | by (blast_tac (!claset addSDs [lemma]) 1); | 
| 923 | 277 | qed "trancl_subset_Sigma"; | 
| 1130 | 278 |