| 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 | 
 | 
| 5067 |      9 | Goal "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))";
 | 
| 0 |     10 | by (rtac bnd_monoI 1);
 | 
|  |     11 | by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2));
 | 
| 2929 |     12 | by (Blast_tac 1);
 | 
| 760 |     13 | qed "rtrancl_bnd_mono";
 | 
| 0 |     14 | 
 | 
| 5321 |     15 | Goalw [rtrancl_def] "r<=s ==> r^* <= s^*";
 | 
| 0 |     16 | by (rtac lfp_mono 1);
 | 
| 5321 |     17 | by (REPEAT (ares_tac [rtrancl_bnd_mono, subset_refl, id_mono,
 | 
| 8318 |     18 | 		      comp_mono, Un_mono, field_mono, Sigma_mono] 1));
 | 
| 760 |     19 | qed "rtrancl_mono";
 | 
| 0 |     20 | 
 | 
|  |     21 | (* r^* = id(field(r)) Un ( r O r^* )    *)
 | 
| 10216 |     22 | bind_thm ("rtrancl_unfold", rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_unfold));
 | 
| 0 |     23 | 
 | 
|  |     24 | (** The relation rtrancl **)
 | 
|  |     25 | 
 | 
| 8318 |     26 | (*  r^* <= field(r) * field(r)  *)
 | 
|  |     27 | bind_thm ("rtrancl_type", rtrancl_def RS def_lfp_subset);
 | 
| 0 |     28 | 
 | 
|  |     29 | (*Reflexivity of rtrancl*)
 | 
| 5321 |     30 | Goal "[| a: field(r) |] ==> <a,a> : r^*";
 | 
| 0 |     31 | by (resolve_tac [rtrancl_unfold RS ssubst] 1);
 | 
| 5321 |     32 | by (etac (idI RS UnI1) 1);
 | 
| 760 |     33 | qed "rtrancl_refl";
 | 
| 0 |     34 | 
 | 
|  |     35 | (*Closure under composition with r  *)
 | 
| 5321 |     36 | Goal "[| <a,b> : r^*;  <b,c> : r |] ==> <a,c> : r^*";
 | 
| 0 |     37 | by (resolve_tac [rtrancl_unfold RS ssubst] 1);
 | 
|  |     38 | by (rtac (compI RS UnI2) 1);
 | 
| 5321 |     39 | by (assume_tac 1);
 | 
|  |     40 | by (assume_tac 1);
 | 
| 760 |     41 | qed "rtrancl_into_rtrancl";
 | 
| 0 |     42 | 
 | 
|  |     43 | (*rtrancl of r contains all pairs in r  *)
 | 
| 5321 |     44 | Goal "<a,b> : r ==> <a,b> : r^*";
 | 
| 0 |     45 | by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1);
 | 
| 5321 |     46 | by (REPEAT (ares_tac [fieldI1] 1));
 | 
| 760 |     47 | qed "r_into_rtrancl";
 | 
| 0 |     48 | 
 | 
|  |     49 | (*The premise ensures that r consists entirely of pairs*)
 | 
| 5321 |     50 | Goal "r <= Sigma(A,B) ==> r <= r^*";
 | 
| 4091 |     51 | by (blast_tac (claset() addIs [r_into_rtrancl]) 1);
 | 
| 760 |     52 | qed "r_subset_rtrancl";
 | 
| 0 |     53 | 
 | 
| 5067 |     54 | Goal "field(r^*) = field(r)";
 | 
| 4091 |     55 | by (blast_tac (claset() addIs [r_into_rtrancl] 
 | 
| 1461 |     56 |                     addSDs [rtrancl_type RS subsetD]) 1);
 | 
| 760 |     57 | qed "rtrancl_field";
 | 
| 0 |     58 | 
 | 
|  |     59 | 
 | 
|  |     60 | (** standard induction rule **)
 | 
|  |     61 | 
 | 
| 5321 |     62 | val major::prems = Goal
 | 
| 0 |     63 |   "[| <a,b> : r^*; \
 | 
|  |     64 | \     !!x. x: field(r) ==> P(<x,x>); \
 | 
|  |     65 | \     !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |]  ==>  P(<x,z>) |] \
 | 
|  |     66 | \  ==>  P(<a,b>)";
 | 
|  |     67 | by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1);
 | 
| 4091 |     68 | by (blast_tac (claset() addIs prems) 1);
 | 
| 760 |     69 | qed "rtrancl_full_induct";
 | 
| 0 |     70 | 
 | 
|  |     71 | (*nice induction rule.
 | 
|  |     72 |   Tried adding the typing hypotheses y,z:field(r), but these
 | 
|  |     73 |   caused expensive case splits!*)
 | 
| 5321 |     74 | val major::prems = Goal
 | 
| 1461 |     75 |   "[| <a,b> : r^*;                                              \
 | 
|  |     76 | \     P(a);                                                     \
 | 
|  |     77 | \     !!y z.[| <a,y> : r^*;  <y,z> : r;  P(y) |] ==> P(z)       \
 | 
| 0 |     78 | \  |] ==> P(b)";
 | 
|  |     79 | (*by induction on this formula*)
 | 
|  |     80 | by (subgoal_tac "ALL y. <a,b> = <a,y> --> P(y)" 1);
 | 
|  |     81 | (*now solve first subgoal: this formula is sufficient*)
 | 
|  |     82 | by (EVERY1 [etac (spec RS mp), rtac refl]);
 | 
|  |     83 | (*now do the induction*)
 | 
|  |     84 | by (resolve_tac [major RS rtrancl_full_induct] 1);
 | 
| 4091 |     85 | by (ALLGOALS (blast_tac (claset() addIs prems)));
 | 
| 760 |     86 | qed "rtrancl_induct";
 | 
| 0 |     87 | 
 | 
|  |     88 | (*transitivity of transitive closure!! -- by induction.*)
 | 
| 5067 |     89 | Goalw [trans_def] "trans(r^*)";
 | 
| 0 |     90 | by (REPEAT (resolve_tac [allI,impI] 1));
 | 
|  |     91 | by (eres_inst_tac [("b","z")] rtrancl_induct 1);
 | 
|  |     92 | by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1));
 | 
| 760 |     93 | qed "trans_rtrancl";
 | 
| 0 |     94 | 
 | 
| 8318 |     95 | bind_thm ("rtrancl_trans", trans_rtrancl RS transD);
 | 
|  |     96 | 
 | 
| 0 |     97 | (*elimination of rtrancl -- by induction on a special formula*)
 | 
| 5321 |     98 | val major::prems = Goal
 | 
| 1461 |     99 |     "[| <a,b> : r^*;  (a=b) ==> P;                       \
 | 
|  |    100 | \       !!y.[| <a,y> : r^*;   <y,b> : r |] ==> P |]      \
 | 
| 0 |    101 | \    ==> P";
 | 
|  |    102 | by (subgoal_tac "a = b  | (EX y. <a,y> : r^* & <y,b> : r)" 1);
 | 
|  |    103 | (*see HOL/trancl*)
 | 
|  |    104 | by (rtac (major RS rtrancl_induct) 2);
 | 
| 4091 |    105 | by (ALLGOALS (fast_tac (claset() addSEs prems)));
 | 
| 760 |    106 | qed "rtranclE";
 | 
| 0 |    107 | 
 | 
|  |    108 | 
 | 
|  |    109 | (**** The relation trancl ****)
 | 
|  |    110 | 
 | 
|  |    111 | (*Transitivity of r^+ is proved by transitivity of r^*  *)
 | 
| 5067 |    112 | Goalw [trans_def,trancl_def] "trans(r^+)";
 | 
| 4091 |    113 | by (blast_tac (claset() addIs [rtrancl_into_rtrancl RS 
 | 
| 3016 |    114 | 			      (trans_rtrancl RS transD RS compI)]) 1);
 | 
| 760 |    115 | qed "trans_trancl";
 | 
| 0 |    116 | 
 | 
| 8318 |    117 | bind_thm ("trancl_trans", trans_trancl RS transD);
 | 
|  |    118 | 
 | 
| 0 |    119 | (** Conversions between trancl and rtrancl **)
 | 
|  |    120 | 
 | 
| 5137 |    121 | Goalw [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*";
 | 
| 4091 |    122 | by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1);
 | 
| 760 |    123 | qed "trancl_into_rtrancl";
 | 
| 0 |    124 | 
 | 
|  |    125 | (*r^+ contains all pairs in r  *)
 | 
| 5137 |    126 | Goalw [trancl_def] "<a,b> : r ==> <a,b> : r^+";
 | 
| 4091 |    127 | by (blast_tac (claset() addSIs [rtrancl_refl]) 1);
 | 
| 760 |    128 | qed "r_into_trancl";
 | 
| 0 |    129 | 
 | 
|  |    130 | (*The premise ensures that r consists entirely of pairs*)
 | 
| 5137 |    131 | Goal "r <= Sigma(A,B) ==> r <= r^+";
 | 
| 4091 |    132 | by (blast_tac (claset() addIs [r_into_trancl]) 1);
 | 
| 760 |    133 | qed "r_subset_trancl";
 | 
| 0 |    134 | 
 | 
|  |    135 | (*intro rule by definition: from r^* and r  *)
 | 
| 9211 |    136 | Goalw [trancl_def] "[| <a,b> : r^*;  <b,c> : r |]   ==>  <a,c> : r^+";
 | 
| 3016 |    137 | by (Blast_tac 1);
 | 
| 760 |    138 | qed "rtrancl_into_trancl1";
 | 
| 0 |    139 | 
 | 
|  |    140 | (*intro rule from r and r^*  *)
 | 
| 9907 |    141 | val prems = goal (the_context ())
 | 
| 0 |    142 |     "[| <a,b> : r;  <b,c> : r^* |]   ==>  <a,c> : r^+";
 | 
|  |    143 | by (resolve_tac (prems RL [rtrancl_induct]) 1);
 | 
|  |    144 | by (resolve_tac (prems RL [r_into_trancl]) 1);
 | 
| 8318 |    145 | by (etac trancl_trans 1);
 | 
| 0 |    146 | by (etac r_into_trancl 1);
 | 
| 760 |    147 | qed "rtrancl_into_trancl2";
 | 
| 0 |    148 | 
 | 
|  |    149 | (*Nice induction rule for trancl*)
 | 
| 5321 |    150 | val major::prems = Goal
 | 
| 1461 |    151 |   "[| <a,b> : r^+;                                      \
 | 
|  |    152 | \     !!y.  [| <a,y> : r |] ==> P(y);                   \
 | 
|  |    153 | \     !!y z.[| <a,y> : r^+;  <y,z> : r;  P(y) |] ==> P(z)       \
 | 
| 0 |    154 | \  |] ==> P(b)";
 | 
|  |    155 | by (rtac (rewrite_rule [trancl_def] major  RS  compEpair) 1);
 | 
|  |    156 | (*by induction on this formula*)
 | 
|  |    157 | by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1);
 | 
|  |    158 | (*now solve first subgoal: this formula is sufficient*)
 | 
| 2929 |    159 | by (Blast_tac 1);
 | 
| 0 |    160 | by (etac rtrancl_induct 1);
 | 
| 4091 |    161 | by (ALLGOALS (fast_tac (claset() addIs (rtrancl_into_trancl1::prems))));
 | 
| 760 |    162 | qed "trancl_induct";
 | 
| 0 |    163 | 
 | 
|  |    164 | (*elimination of r^+ -- NOT an induction rule*)
 | 
| 5321 |    165 | val major::prems = Goal
 | 
| 0 |    166 |     "[| <a,b> : r^+;  \
 | 
|  |    167 | \       <a,b> : r ==> P; \
 | 
| 1461 |    168 | \       !!y.[| <a,y> : r^+; <y,b> : r |] ==> P  \
 | 
| 0 |    169 | \    |] ==> P";
 | 
|  |    170 | by (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+  &  <y,b> : r)" 1);
 | 
| 4091 |    171 | by (fast_tac (claset() addIs prems) 1);
 | 
| 0 |    172 | by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
 | 
|  |    173 | by (etac rtranclE 1);
 | 
| 4091 |    174 | by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_trancl1])));
 | 
| 760 |    175 | qed "tranclE";
 | 
| 0 |    176 | 
 | 
| 5067 |    177 | Goalw [trancl_def] "r^+ <= field(r)*field(r)";
 | 
| 4091 |    178 | by (blast_tac (claset() addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);
 | 
| 760 |    179 | qed "trancl_type";
 | 
| 0 |    180 | 
 | 
| 5321 |    181 | Goalw [trancl_def] "r<=s ==> r^+ <= s^+";
 | 
|  |    182 | by (REPEAT (ares_tac [comp_mono, rtrancl_mono] 1));
 | 
| 760 |    183 | qed "trancl_mono";
 | 
| 0 |    184 | 
 | 
| 8318 |    185 | (** Suggested by Sidi Ould Ehmety **)
 | 
|  |    186 | 
 | 
|  |    187 | Goal "(r^*)^* = r^*";
 | 
|  |    188 | by (rtac equalityI 1);
 | 
|  |    189 | by Auto_tac;
 | 
|  |    190 | by (ALLGOALS (forward_tac [impOfSubs rtrancl_type]));
 | 
|  |    191 | by (ALLGOALS Clarify_tac);
 | 
|  |    192 | by (etac r_into_rtrancl 2);
 | 
|  |    193 | by (etac rtrancl_induct 1);
 | 
|  |    194 | by (asm_full_simp_tac (simpset() addsimps [rtrancl_refl, rtrancl_field]) 1);
 | 
|  |    195 | by (blast_tac (claset() addIs [rtrancl_trans]) 1);
 | 
|  |    196 | qed "rtrancl_idemp";
 | 
|  |    197 | Addsimps [rtrancl_idemp];
 | 
|  |    198 | 
 | 
|  |    199 | Goal "[| R <= S; S <= R^* |] ==> S^* = R^*";
 | 
|  |    200 | by (dtac rtrancl_mono 1);
 | 
|  |    201 | by (dtac rtrancl_mono 1);
 | 
|  |    202 | by (ALLGOALS Asm_full_simp_tac);
 | 
|  |    203 | by (Blast_tac 1);
 | 
|  |    204 | qed "rtrancl_subset";
 | 
|  |    205 | 
 | 
|  |    206 | Goal "[| r<= Sigma(A,B); s<=Sigma(C,D) |] ==> (r^* Un s^*)^* = (r Un s)^*";
 | 
|  |    207 | by (rtac rtrancl_subset 1);
 | 
|  |    208 | by (blast_tac (claset() addDs [r_subset_rtrancl]) 1);
 | 
|  |    209 | by (blast_tac (claset() addIs [rtrancl_mono RS subsetD]) 1);
 | 
|  |    210 | qed "rtrancl_Un_rtrancl";
 | 
|  |    211 | 
 | 
|  |    212 | (** "converse" laws by Sidi Ould Ehmety **)
 | 
|  |    213 | 
 | 
|  |    214 | Goal "<x,y>:converse(r)^* ==> <x,y>:converse(r^*)";
 | 
|  |    215 | by (rtac converseI 1);
 | 
|  |    216 | by (forward_tac [rtrancl_type RS subsetD] 1);
 | 
|  |    217 | by (etac rtrancl_induct 1);
 | 
|  |    218 | by (blast_tac (claset() addIs [rtrancl_refl]) 1);
 | 
|  |    219 | by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1);
 | 
|  |    220 | qed "rtrancl_converseD";
 | 
|  |    221 | 
 | 
|  |    222 | Goal "<x,y>:converse(r^*) ==> <x,y>:converse(r)^*";
 | 
|  |    223 | by (dtac converseD 1);
 | 
|  |    224 | by (forward_tac [rtrancl_type RS subsetD] 1);
 | 
|  |    225 | by (etac rtrancl_induct 1);
 | 
|  |    226 | by (blast_tac (claset() addIs [rtrancl_refl]) 1);
 | 
|  |    227 | by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1);
 | 
|  |    228 | qed "rtrancl_converseI";
 | 
|  |    229 | 
 | 
|  |    230 | Goal "converse(r)^* = converse(r^*)";
 | 
|  |    231 | by (safe_tac (claset() addSIs [equalityI]));
 | 
|  |    232 | by (forward_tac [rtrancl_type RS subsetD] 1);
 | 
|  |    233 | by (safe_tac (claset() addSDs [rtrancl_converseD] addSIs [rtrancl_converseI]));
 | 
|  |    234 | qed "rtrancl_converse";
 |