| author | wenzelm | 
| Fri, 22 Oct 1999 21:49:33 +0200 | |
| changeset 7924 | 5fee69b1f5fe | 
| parent 5321 | f8848433d240 | 
| child 8318 | 54d69141a17f | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/trancl.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 2929 | 6 | Transitive closure of a relation | 
| 0 | 7 | *) | 
| 8 | ||
| 9 | open Trancl; | |
| 10 | ||
| 5067 | 11 | Goal "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"; | 
| 0 | 12 | by (rtac bnd_monoI 1); | 
| 13 | by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2)); | |
| 2929 | 14 | by (Blast_tac 1); | 
| 760 | 15 | qed "rtrancl_bnd_mono"; | 
| 0 | 16 | |
| 5321 | 17 | Goalw [rtrancl_def] "r<=s ==> r^* <= s^*"; | 
| 0 | 18 | by (rtac lfp_mono 1); | 
| 5321 | 19 | by (REPEAT (ares_tac [rtrancl_bnd_mono, subset_refl, id_mono, | 
| 1461 | 20 | comp_mono, Un_mono, field_mono, Sigma_mono] 1)); | 
| 760 | 21 | qed "rtrancl_mono"; | 
| 0 | 22 | |
| 23 | (* r^* = id(field(r)) Un ( r O r^* ) *) | |
| 24 | val rtrancl_unfold = rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_Tarski); | |
| 25 | ||
| 26 | (** The relation rtrancl **) | |
| 27 | ||
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 28 | bind_thm ("rtrancl_type", (rtrancl_def RS def_lfp_subset));
 | 
| 0 | 29 | |
| 30 | (*Reflexivity of rtrancl*) | |
| 5321 | 31 | Goal "[| a: field(r) |] ==> <a,a> : r^*"; | 
| 0 | 32 | by (resolve_tac [rtrancl_unfold RS ssubst] 1); | 
| 5321 | 33 | by (etac (idI RS UnI1) 1); | 
| 760 | 34 | qed "rtrancl_refl"; | 
| 0 | 35 | |
| 36 | (*Closure under composition with r *) | |
| 5321 | 37 | Goal "[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^*"; | 
| 0 | 38 | by (resolve_tac [rtrancl_unfold RS ssubst] 1); | 
| 39 | by (rtac (compI RS UnI2) 1); | |
| 5321 | 40 | by (assume_tac 1); | 
| 41 | by (assume_tac 1); | |
| 760 | 42 | qed "rtrancl_into_rtrancl"; | 
| 0 | 43 | |
| 44 | (*rtrancl of r contains all pairs in r *) | |
| 5321 | 45 | Goal "<a,b> : r ==> <a,b> : r^*"; | 
| 0 | 46 | by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1); | 
| 5321 | 47 | by (REPEAT (ares_tac [fieldI1] 1)); | 
| 760 | 48 | qed "r_into_rtrancl"; | 
| 0 | 49 | |
| 50 | (*The premise ensures that r consists entirely of pairs*) | |
| 5321 | 51 | Goal "r <= Sigma(A,B) ==> r <= r^*"; | 
| 4091 | 52 | by (blast_tac (claset() addIs [r_into_rtrancl]) 1); | 
| 760 | 53 | qed "r_subset_rtrancl"; | 
| 0 | 54 | |
| 5067 | 55 | Goal "field(r^*) = field(r)"; | 
| 4091 | 56 | by (blast_tac (claset() addIs [r_into_rtrancl] | 
| 1461 | 57 | addSDs [rtrancl_type RS subsetD]) 1); | 
| 760 | 58 | qed "rtrancl_field"; | 
| 0 | 59 | |
| 60 | ||
| 61 | (** standard induction rule **) | |
| 62 | ||
| 5321 | 63 | val major::prems = Goal | 
| 0 | 64 | "[| <a,b> : r^*; \ | 
| 65 | \ !!x. x: field(r) ==> P(<x,x>); \ | |
| 66 | \ !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |] \ | |
| 67 | \ ==> P(<a,b>)"; | |
| 68 | by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1); | |
| 4091 | 69 | by (blast_tac (claset() addIs prems) 1); | 
| 760 | 70 | qed "rtrancl_full_induct"; | 
| 0 | 71 | |
| 72 | (*nice induction rule. | |
| 73 | Tried adding the typing hypotheses y,z:field(r), but these | |
| 74 | caused expensive case splits!*) | |
| 5321 | 75 | val major::prems = Goal | 
| 1461 | 76 | "[| <a,b> : r^*; \ | 
| 77 | \ P(a); \ | |
| 78 | \ !!y z.[| <a,y> : r^*; <y,z> : r; P(y) |] ==> P(z) \ | |
| 0 | 79 | \ |] ==> P(b)"; | 
| 80 | (*by induction on this formula*) | |
| 81 | by (subgoal_tac "ALL y. <a,b> = <a,y> --> P(y)" 1); | |
| 82 | (*now solve first subgoal: this formula is sufficient*) | |
| 83 | by (EVERY1 [etac (spec RS mp), rtac refl]); | |
| 84 | (*now do the induction*) | |
| 85 | by (resolve_tac [major RS rtrancl_full_induct] 1); | |
| 4091 | 86 | by (ALLGOALS (blast_tac (claset() addIs prems))); | 
| 760 | 87 | qed "rtrancl_induct"; | 
| 0 | 88 | |
| 89 | (*transitivity of transitive closure!! -- by induction.*) | |
| 5067 | 90 | Goalw [trans_def] "trans(r^*)"; | 
| 0 | 91 | by (REPEAT (resolve_tac [allI,impI] 1)); | 
| 92 | by (eres_inst_tac [("b","z")] rtrancl_induct 1);
 | |
| 93 | by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); | |
| 760 | 94 | qed "trans_rtrancl"; | 
| 0 | 95 | |
| 96 | (*elimination of rtrancl -- by induction on a special formula*) | |
| 5321 | 97 | val major::prems = Goal | 
| 1461 | 98 | "[| <a,b> : r^*; (a=b) ==> P; \ | 
| 99 | \ !!y.[| <a,y> : r^*; <y,b> : r |] ==> P |] \ | |
| 0 | 100 | \ ==> P"; | 
| 101 | by (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r)" 1); | |
| 102 | (*see HOL/trancl*) | |
| 103 | by (rtac (major RS rtrancl_induct) 2); | |
| 4091 | 104 | by (ALLGOALS (fast_tac (claset() addSEs prems))); | 
| 760 | 105 | qed "rtranclE"; | 
| 0 | 106 | |
| 107 | ||
| 108 | (**** The relation trancl ****) | |
| 109 | ||
| 110 | (*Transitivity of r^+ is proved by transitivity of r^* *) | |
| 5067 | 111 | Goalw [trans_def,trancl_def] "trans(r^+)"; | 
| 4091 | 112 | by (blast_tac (claset() addIs [rtrancl_into_rtrancl RS | 
| 3016 | 113 | (trans_rtrancl RS transD RS compI)]) 1); | 
| 760 | 114 | qed "trans_trancl"; | 
| 0 | 115 | |
| 116 | (** Conversions between trancl and rtrancl **) | |
| 117 | ||
| 5137 | 118 | Goalw [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*"; | 
| 4091 | 119 | by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); | 
| 760 | 120 | qed "trancl_into_rtrancl"; | 
| 0 | 121 | |
| 122 | (*r^+ contains all pairs in r *) | |
| 5137 | 123 | Goalw [trancl_def] "<a,b> : r ==> <a,b> : r^+"; | 
| 4091 | 124 | by (blast_tac (claset() addSIs [rtrancl_refl]) 1); | 
| 760 | 125 | qed "r_into_trancl"; | 
| 0 | 126 | |
| 127 | (*The premise ensures that r consists entirely of pairs*) | |
| 5137 | 128 | Goal "r <= Sigma(A,B) ==> r <= r^+"; | 
| 4091 | 129 | by (blast_tac (claset() addIs [r_into_trancl]) 1); | 
| 760 | 130 | qed "r_subset_trancl"; | 
| 0 | 131 | |
| 132 | (*intro rule by definition: from r^* and r *) | |
| 5067 | 133 | Goalw [trancl_def] | 
| 3016 | 134 | "!!r. [| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+"; | 
| 135 | by (Blast_tac 1); | |
| 760 | 136 | qed "rtrancl_into_trancl1"; | 
| 0 | 137 | |
| 138 | (*intro rule from r and r^* *) | |
| 139 | val prems = goal Trancl.thy | |
| 140 | "[| <a,b> : r; <b,c> : r^* |] ==> <a,c> : r^+"; | |
| 141 | by (resolve_tac (prems RL [rtrancl_induct]) 1); | |
| 142 | by (resolve_tac (prems RL [r_into_trancl]) 1); | |
| 143 | by (etac (trans_trancl RS transD) 1); | |
| 144 | by (etac r_into_trancl 1); | |
| 760 | 145 | qed "rtrancl_into_trancl2"; | 
| 0 | 146 | |
| 147 | (*Nice induction rule for trancl*) | |
| 5321 | 148 | val major::prems = Goal | 
| 1461 | 149 | "[| <a,b> : r^+; \ | 
| 150 | \ !!y. [| <a,y> : r |] ==> P(y); \ | |
| 151 | \ !!y z.[| <a,y> : r^+; <y,z> : r; P(y) |] ==> P(z) \ | |
| 0 | 152 | \ |] ==> P(b)"; | 
| 153 | by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); | |
| 154 | (*by induction on this formula*) | |
| 155 | by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1); | |
| 156 | (*now solve first subgoal: this formula is sufficient*) | |
| 2929 | 157 | by (Blast_tac 1); | 
| 0 | 158 | by (etac rtrancl_induct 1); | 
| 4091 | 159 | by (ALLGOALS (fast_tac (claset() addIs (rtrancl_into_trancl1::prems)))); | 
| 760 | 160 | qed "trancl_induct"; | 
| 0 | 161 | |
| 162 | (*elimination of r^+ -- NOT an induction rule*) | |
| 5321 | 163 | val major::prems = Goal | 
| 0 | 164 | "[| <a,b> : r^+; \ | 
| 165 | \ <a,b> : r ==> P; \ | |
| 1461 | 166 | \ !!y.[| <a,y> : r^+; <y,b> : r |] ==> P \ | 
| 0 | 167 | \ |] ==> P"; | 
| 168 | by (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r)" 1); | |
| 4091 | 169 | by (fast_tac (claset() addIs prems) 1); | 
| 0 | 170 | by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); | 
| 171 | by (etac rtranclE 1); | |
| 4091 | 172 | by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_trancl1]))); | 
| 760 | 173 | qed "tranclE"; | 
| 0 | 174 | |
| 5067 | 175 | Goalw [trancl_def] "r^+ <= field(r)*field(r)"; | 
| 4091 | 176 | by (blast_tac (claset() addEs [rtrancl_type RS subsetD RS SigmaE2]) 1); | 
| 760 | 177 | qed "trancl_type"; | 
| 0 | 178 | |
| 5321 | 179 | Goalw [trancl_def] "r<=s ==> r^+ <= s^+"; | 
| 180 | by (REPEAT (ares_tac [comp_mono, rtrancl_mono] 1)); | |
| 760 | 181 | qed "trancl_mono"; | 
| 0 | 182 |