| 13356 |      1 | (*  Title:      ZF/Trancl.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1478 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
| 13356 |      8 | header{*Relations: Their General Properties and Transitive Closure*}
 | 
|  |      9 | 
 | 
| 16417 |     10 | theory Trancl imports Fixedpt Perm begin
 | 
| 13239 |     11 | 
 | 
| 24893 |     12 | definition
 | 
|  |     13 |   refl     :: "[i,i]=>o"  where
 | 
| 13239 |     14 |     "refl(A,r) == (ALL x: A. <x,x> : r)"
 | 
|  |     15 | 
 | 
| 24893 |     16 | definition
 | 
|  |     17 |   irrefl   :: "[i,i]=>o"  where
 | 
| 13239 |     18 |     "irrefl(A,r) == ALL x: A. <x,x> ~: r"
 | 
|  |     19 | 
 | 
| 24893 |     20 | definition
 | 
|  |     21 |   sym      :: "i=>o"  where
 | 
| 13239 |     22 |     "sym(r) == ALL x y. <x,y>: r --> <y,x>: r"
 | 
|  |     23 | 
 | 
| 24893 |     24 | definition
 | 
|  |     25 |   asym     :: "i=>o"  where
 | 
| 13239 |     26 |     "asym(r) == ALL x y. <x,y>:r --> ~ <y,x>:r"
 | 
|  |     27 | 
 | 
| 24893 |     28 | definition
 | 
|  |     29 |   antisym  :: "i=>o"  where
 | 
| 13239 |     30 |     "antisym(r) == ALL x y.<x,y>:r --> <y,x>:r --> x=y"
 | 
|  |     31 | 
 | 
| 24893 |     32 | definition
 | 
|  |     33 |   trans    :: "i=>o"  where
 | 
| 13239 |     34 |     "trans(r) == ALL x y z. <x,y>: r --> <y,z>: r --> <x,z>: r"
 | 
|  |     35 | 
 | 
| 24893 |     36 | definition
 | 
|  |     37 |   trans_on :: "[i,i]=>o"  ("trans[_]'(_')")  where
 | 
| 13239 |     38 |     "trans[A](r) == ALL x:A. ALL y:A. ALL z:A.       
 | 
|  |     39 |                           <x,y>: r --> <y,z>: r --> <x,z>: r"
 | 
|  |     40 | 
 | 
| 24893 |     41 | definition
 | 
|  |     42 |   rtrancl :: "i=>i"  ("(_^*)" [100] 100)  (*refl/transitive closure*)  where
 | 
| 13239 |     43 |     "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))"
 | 
|  |     44 | 
 | 
| 24893 |     45 | definition
 | 
|  |     46 |   trancl  :: "i=>i"  ("(_^+)" [100] 100)  (*transitive closure*)  where
 | 
| 13239 |     47 |     "r^+ == r O r^*"
 | 
|  |     48 | 
 | 
| 24893 |     49 | definition
 | 
|  |     50 |   equiv    :: "[i,i]=>o"  where
 | 
| 14653 |     51 |     "equiv(A,r) == r <= A*A & refl(A,r) & sym(r) & trans(r)"
 | 
|  |     52 | 
 | 
|  |     53 | 
 | 
| 13239 |     54 | subsection{*General properties of relations*}
 | 
|  |     55 | 
 | 
|  |     56 | subsubsection{*irreflexivity*}
 | 
|  |     57 | 
 | 
|  |     58 | lemma irreflI:
 | 
|  |     59 |     "[| !!x. x:A ==> <x,x> ~: r |] ==> irrefl(A,r)"
 | 
| 13240 |     60 | by (simp add: irrefl_def) 
 | 
| 13239 |     61 | 
 | 
| 13534 |     62 | lemma irreflE: "[| irrefl(A,r);  x:A |] ==>  <x,x> ~: r"
 | 
| 13240 |     63 | by (simp add: irrefl_def)
 | 
| 13239 |     64 | 
 | 
|  |     65 | subsubsection{*symmetry*}
 | 
|  |     66 | 
 | 
|  |     67 | lemma symI:
 | 
|  |     68 |      "[| !!x y.<x,y>: r ==> <y,x>: r |] ==> sym(r)"
 | 
| 13356 |     69 | by (unfold sym_def, blast) 
 | 
| 13239 |     70 | 
 | 
| 13534 |     71 | lemma symE: "[| sym(r); <x,y>: r |]  ==>  <y,x>: r"
 | 
| 13240 |     72 | by (unfold sym_def, blast)
 | 
| 13239 |     73 | 
 | 
|  |     74 | subsubsection{*antisymmetry*}
 | 
|  |     75 | 
 | 
|  |     76 | lemma antisymI:
 | 
|  |     77 |      "[| !!x y.[| <x,y>: r;  <y,x>: r |] ==> x=y |] ==> antisym(r)"
 | 
| 13356 |     78 | by (simp add: antisym_def, blast) 
 | 
| 13239 |     79 | 
 | 
|  |     80 | lemma antisymE: "[| antisym(r); <x,y>: r;  <y,x>: r |]  ==>  x=y"
 | 
| 13240 |     81 | by (simp add: antisym_def, blast)
 | 
| 13239 |     82 | 
 | 
|  |     83 | subsubsection{*transitivity*}
 | 
|  |     84 | 
 | 
|  |     85 | lemma transD: "[| trans(r);  <a,b>:r;  <b,c>:r |] ==> <a,c>:r"
 | 
| 13240 |     86 | by (unfold trans_def, blast)
 | 
| 13239 |     87 | 
 | 
|  |     88 | lemma trans_onD: 
 | 
|  |     89 |     "[| trans[A](r);  <a,b>:r;  <b,c>:r;  a:A;  b:A;  c:A |] ==> <a,c>:r"
 | 
| 13243 |     90 | by (unfold trans_on_def, blast)
 | 
|  |     91 | 
 | 
|  |     92 | lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
 | 
|  |     93 | by (unfold trans_def trans_on_def, blast)
 | 
| 13239 |     94 | 
 | 
| 13248 |     95 | lemma trans_on_imp_trans: "[|trans[A](r); r <= A*A|] ==> trans(r)";
 | 
|  |     96 | by (simp add: trans_on_def trans_def, blast)
 | 
|  |     97 | 
 | 
|  |     98 | 
 | 
| 13239 |     99 | subsection{*Transitive closure of a relation*}
 | 
|  |    100 | 
 | 
|  |    101 | lemma rtrancl_bnd_mono:
 | 
|  |    102 |      "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"
 | 
| 13248 |    103 | by (rule bnd_monoI, blast+)
 | 
| 13239 |    104 | 
 | 
|  |    105 | lemma rtrancl_mono: "r<=s ==> r^* <= s^*"
 | 
|  |    106 | apply (unfold rtrancl_def)
 | 
|  |    107 | apply (rule lfp_mono)
 | 
|  |    108 | apply (rule rtrancl_bnd_mono)+
 | 
| 13240 |    109 | apply blast 
 | 
| 13239 |    110 | done
 | 
|  |    111 | 
 | 
|  |    112 | (* r^* = id(field(r)) Un ( r O r^* )    *)
 | 
|  |    113 | lemmas rtrancl_unfold =
 | 
|  |    114 |      rtrancl_bnd_mono [THEN rtrancl_def [THEN def_lfp_unfold], standard]
 | 
|  |    115 | 
 | 
|  |    116 | (** The relation rtrancl **)
 | 
|  |    117 | 
 | 
|  |    118 | (*  r^* <= field(r) * field(r)  *)
 | 
|  |    119 | lemmas rtrancl_type = rtrancl_def [THEN def_lfp_subset, standard]
 | 
|  |    120 | 
 | 
| 13248 |    121 | lemma relation_rtrancl: "relation(r^*)"
 | 
|  |    122 | apply (simp add: relation_def) 
 | 
|  |    123 | apply (blast dest: rtrancl_type [THEN subsetD]) 
 | 
|  |    124 | done
 | 
|  |    125 | 
 | 
| 13239 |    126 | (*Reflexivity of rtrancl*)
 | 
|  |    127 | lemma rtrancl_refl: "[| a: field(r) |] ==> <a,a> : r^*"
 | 
|  |    128 | apply (rule rtrancl_unfold [THEN ssubst])
 | 
|  |    129 | apply (erule idI [THEN UnI1])
 | 
|  |    130 | done
 | 
|  |    131 | 
 | 
|  |    132 | (*Closure under composition with r  *)
 | 
|  |    133 | lemma rtrancl_into_rtrancl: "[| <a,b> : r^*;  <b,c> : r |] ==> <a,c> : r^*"
 | 
|  |    134 | apply (rule rtrancl_unfold [THEN ssubst])
 | 
| 13269 |    135 | apply (rule compI [THEN UnI2], assumption, assumption)
 | 
| 13239 |    136 | done
 | 
|  |    137 | 
 | 
|  |    138 | (*rtrancl of r contains all pairs in r  *)
 | 
|  |    139 | lemma r_into_rtrancl: "<a,b> : r ==> <a,b> : r^*"
 | 
| 13240 |    140 | by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+)
 | 
| 13239 |    141 | 
 | 
|  |    142 | (*The premise ensures that r consists entirely of pairs*)
 | 
| 13248 |    143 | lemma r_subset_rtrancl: "relation(r) ==> r <= r^*"
 | 
|  |    144 | by (simp add: relation_def, blast intro: r_into_rtrancl)
 | 
| 13239 |    145 | 
 | 
|  |    146 | lemma rtrancl_field: "field(r^*) = field(r)"
 | 
| 13240 |    147 | by (blast intro: r_into_rtrancl dest!: rtrancl_type [THEN subsetD])
 | 
| 13239 |    148 | 
 | 
|  |    149 | 
 | 
|  |    150 | (** standard induction rule **)
 | 
|  |    151 | 
 | 
| 13534 |    152 | lemma rtrancl_full_induct [case_names initial step, consumes 1]:
 | 
| 13239 |    153 |   "[| <a,b> : r^*;  
 | 
|  |    154 |       !!x. x: field(r) ==> P(<x,x>);  
 | 
|  |    155 |       !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |]  ==>  P(<x,z>) |]  
 | 
|  |    156 |    ==>  P(<a,b>)"
 | 
| 13248 |    157 | by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast) 
 | 
| 13239 |    158 | 
 | 
|  |    159 | (*nice induction rule.
 | 
|  |    160 |   Tried adding the typing hypotheses y,z:field(r), but these
 | 
|  |    161 |   caused expensive case splits!*)
 | 
| 13534 |    162 | lemma rtrancl_induct [case_names initial step, induct set: rtrancl]:
 | 
| 13239 |    163 |   "[| <a,b> : r^*;                                               
 | 
|  |    164 |       P(a);                                                      
 | 
|  |    165 |       !!y z.[| <a,y> : r^*;  <y,z> : r;  P(y) |] ==> P(z)        
 | 
|  |    166 |    |] ==> P(b)"
 | 
|  |    167 | (*by induction on this formula*)
 | 
|  |    168 | apply (subgoal_tac "ALL y. <a,b> = <a,y> --> P (y) ")
 | 
|  |    169 | (*now solve first subgoal: this formula is sufficient*)
 | 
|  |    170 | apply (erule spec [THEN mp], rule refl)
 | 
|  |    171 | (*now do the induction*)
 | 
| 13240 |    172 | apply (erule rtrancl_full_induct, blast+)
 | 
| 13239 |    173 | done
 | 
|  |    174 | 
 | 
|  |    175 | (*transitivity of transitive closure!! -- by induction.*)
 | 
|  |    176 | lemma trans_rtrancl: "trans(r^*)"
 | 
|  |    177 | apply (unfold trans_def)
 | 
|  |    178 | apply (intro allI impI)
 | 
| 13784 |    179 | apply (erule_tac b = z in rtrancl_induct, assumption)
 | 
| 13239 |    180 | apply (blast intro: rtrancl_into_rtrancl) 
 | 
|  |    181 | done
 | 
|  |    182 | 
 | 
|  |    183 | lemmas rtrancl_trans = trans_rtrancl [THEN transD, standard]
 | 
|  |    184 | 
 | 
|  |    185 | (*elimination of rtrancl -- by induction on a special formula*)
 | 
|  |    186 | lemma rtranclE:
 | 
|  |    187 |     "[| <a,b> : r^*;  (a=b) ==> P;                        
 | 
|  |    188 |         !!y.[| <a,y> : r^*;   <y,b> : r |] ==> P |]       
 | 
|  |    189 |      ==> P"
 | 
|  |    190 | apply (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r) ")
 | 
|  |    191 | (*see HOL/trancl*)
 | 
| 13240 |    192 | apply blast 
 | 
|  |    193 | apply (erule rtrancl_induct, blast+)
 | 
| 13239 |    194 | done
 | 
|  |    195 | 
 | 
|  |    196 | 
 | 
|  |    197 | (**** The relation trancl ****)
 | 
|  |    198 | 
 | 
|  |    199 | (*Transitivity of r^+ is proved by transitivity of r^*  *)
 | 
|  |    200 | lemma trans_trancl: "trans(r^+)"
 | 
|  |    201 | apply (unfold trans_def trancl_def)
 | 
|  |    202 | apply (blast intro: rtrancl_into_rtrancl
 | 
|  |    203 |                     trans_rtrancl [THEN transD, THEN compI])
 | 
|  |    204 | done
 | 
|  |    205 | 
 | 
| 13243 |    206 | lemmas trans_on_trancl = trans_trancl [THEN trans_imp_trans_on]
 | 
|  |    207 | 
 | 
| 13239 |    208 | lemmas trancl_trans = trans_trancl [THEN transD, standard]
 | 
|  |    209 | 
 | 
|  |    210 | (** Conversions between trancl and rtrancl **)
 | 
| 0 |    211 | 
 | 
| 13239 |    212 | lemma trancl_into_rtrancl: "<a,b> : r^+ ==> <a,b> : r^*"
 | 
|  |    213 | apply (unfold trancl_def)
 | 
|  |    214 | apply (blast intro: rtrancl_into_rtrancl)
 | 
|  |    215 | done
 | 
|  |    216 | 
 | 
|  |    217 | (*r^+ contains all pairs in r  *)
 | 
|  |    218 | lemma r_into_trancl: "<a,b> : r ==> <a,b> : r^+"
 | 
|  |    219 | apply (unfold trancl_def)
 | 
|  |    220 | apply (blast intro!: rtrancl_refl)
 | 
|  |    221 | done
 | 
|  |    222 | 
 | 
|  |    223 | (*The premise ensures that r consists entirely of pairs*)
 | 
| 13248 |    224 | lemma r_subset_trancl: "relation(r) ==> r <= r^+"
 | 
|  |    225 | by (simp add: relation_def, blast intro: r_into_trancl)
 | 
|  |    226 | 
 | 
| 13239 |    227 | 
 | 
|  |    228 | (*intro rule by definition: from r^* and r  *)
 | 
|  |    229 | lemma rtrancl_into_trancl1: "[| <a,b> : r^*;  <b,c> : r |]   ==>  <a,c> : r^+"
 | 
| 13240 |    230 | by (unfold trancl_def, blast)
 | 
| 13239 |    231 | 
 | 
|  |    232 | (*intro rule from r and r^*  *)
 | 
|  |    233 | lemma rtrancl_into_trancl2:
 | 
|  |    234 |     "[| <a,b> : r;  <b,c> : r^* |]   ==>  <a,c> : r^+"
 | 
|  |    235 | apply (erule rtrancl_induct)
 | 
|  |    236 |  apply (erule r_into_trancl)
 | 
|  |    237 | apply (blast intro: r_into_trancl trancl_trans) 
 | 
|  |    238 | done
 | 
|  |    239 | 
 | 
|  |    240 | (*Nice induction rule for trancl*)
 | 
| 13534 |    241 | lemma trancl_induct [case_names initial step, induct set: trancl]:
 | 
| 13239 |    242 |   "[| <a,b> : r^+;                                       
 | 
|  |    243 |       !!y.  [| <a,y> : r |] ==> P(y);                    
 | 
|  |    244 |       !!y z.[| <a,y> : r^+;  <y,z> : r;  P(y) |] ==> P(z)        
 | 
|  |    245 |    |] ==> P(b)"
 | 
|  |    246 | apply (rule compEpair)
 | 
|  |    247 | apply (unfold trancl_def, assumption)
 | 
|  |    248 | (*by induction on this formula*)
 | 
|  |    249 | apply (subgoal_tac "ALL z. <y,z> : r --> P (z) ")
 | 
|  |    250 | (*now solve first subgoal: this formula is sufficient*)
 | 
|  |    251 |  apply blast
 | 
|  |    252 | apply (erule rtrancl_induct)
 | 
|  |    253 | apply (blast intro: rtrancl_into_trancl1)+
 | 
|  |    254 | done
 | 
|  |    255 | 
 | 
|  |    256 | (*elimination of r^+ -- NOT an induction rule*)
 | 
|  |    257 | lemma tranclE:
 | 
|  |    258 |     "[| <a,b> : r^+;   
 | 
|  |    259 |         <a,b> : r ==> P;  
 | 
|  |    260 |         !!y.[| <a,y> : r^+; <y,b> : r |] ==> P   
 | 
|  |    261 |      |] ==> P"
 | 
|  |    262 | apply (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r) ")
 | 
| 13240 |    263 | apply blast 
 | 
| 13239 |    264 | apply (rule compEpair)
 | 
|  |    265 | apply (unfold trancl_def, assumption)
 | 
|  |    266 | apply (erule rtranclE)
 | 
|  |    267 | apply (blast intro: rtrancl_into_trancl1)+
 | 
|  |    268 | done
 | 
|  |    269 | 
 | 
|  |    270 | lemma trancl_type: "r^+ <= field(r)*field(r)"
 | 
|  |    271 | apply (unfold trancl_def)
 | 
|  |    272 | apply (blast elim: rtrancl_type [THEN subsetD, THEN SigmaE2])
 | 
|  |    273 | done
 | 
|  |    274 | 
 | 
| 13248 |    275 | lemma relation_trancl: "relation(r^+)"
 | 
|  |    276 | apply (simp add: relation_def) 
 | 
|  |    277 | apply (blast dest: trancl_type [THEN subsetD]) 
 | 
|  |    278 | done
 | 
|  |    279 | 
 | 
| 13243 |    280 | lemma trancl_subset_times: "r \<subseteq> A * A ==> r^+ \<subseteq> A * A"
 | 
|  |    281 | by (insert trancl_type [of r], blast)
 | 
|  |    282 | 
 | 
| 13239 |    283 | lemma trancl_mono: "r<=s ==> r^+ <= s^+"
 | 
|  |    284 | by (unfold trancl_def, intro comp_mono rtrancl_mono)
 | 
|  |    285 | 
 | 
| 13248 |    286 | lemma trancl_eq_r: "[|relation(r); trans(r)|] ==> r^+ = r"
 | 
|  |    287 | apply (rule equalityI)
 | 
|  |    288 |  prefer 2 apply (erule r_subset_trancl, clarify) 
 | 
|  |    289 | apply (frule trancl_type [THEN subsetD], clarify) 
 | 
|  |    290 | apply (erule trancl_induct, assumption)
 | 
|  |    291 | apply (blast dest: transD) 
 | 
|  |    292 | done
 | 
|  |    293 | 
 | 
| 13239 |    294 | 
 | 
|  |    295 | (** Suggested by Sidi Ould Ehmety **)
 | 
|  |    296 | 
 | 
|  |    297 | lemma rtrancl_idemp [simp]: "(r^*)^* = r^*"
 | 
| 13240 |    298 | apply (rule equalityI, auto)
 | 
| 13239 |    299 |  prefer 2
 | 
|  |    300 |  apply (frule rtrancl_type [THEN subsetD])
 | 
| 13240 |    301 |  apply (blast intro: r_into_rtrancl ) 
 | 
| 13239 |    302 | txt{*converse direction*}
 | 
| 13240 |    303 | apply (frule rtrancl_type [THEN subsetD], clarify) 
 | 
| 13239 |    304 | apply (erule rtrancl_induct)
 | 
|  |    305 | apply (simp add: rtrancl_refl rtrancl_field)
 | 
|  |    306 | apply (blast intro: rtrancl_trans)
 | 
|  |    307 | done
 | 
|  |    308 | 
 | 
|  |    309 | lemma rtrancl_subset: "[| R <= S; S <= R^* |] ==> S^* = R^*"
 | 
|  |    310 | apply (drule rtrancl_mono)
 | 
| 13269 |    311 | apply (drule rtrancl_mono, simp_all, blast)
 | 
| 13239 |    312 | done
 | 
|  |    313 | 
 | 
|  |    314 | lemma rtrancl_Un_rtrancl:
 | 
| 13248 |    315 |      "[| relation(r); relation(s) |] ==> (r^* Un s^*)^* = (r Un s)^*"
 | 
| 13239 |    316 | apply (rule rtrancl_subset)
 | 
|  |    317 | apply (blast dest: r_subset_rtrancl)
 | 
|  |    318 | apply (blast intro: rtrancl_mono [THEN subsetD])
 | 
|  |    319 | done
 | 
|  |    320 | 
 | 
|  |    321 | (*** "converse" laws by Sidi Ould Ehmety ***)
 | 
|  |    322 | 
 | 
|  |    323 | (** rtrancl **)
 | 
|  |    324 | 
 | 
|  |    325 | lemma rtrancl_converseD: "<x,y>:converse(r)^* ==> <x,y>:converse(r^*)"
 | 
|  |    326 | apply (rule converseI)
 | 
|  |    327 | apply (frule rtrancl_type [THEN subsetD])
 | 
|  |    328 | apply (erule rtrancl_induct)
 | 
|  |    329 | apply (blast intro: rtrancl_refl)
 | 
|  |    330 | apply (blast intro: r_into_rtrancl rtrancl_trans)
 | 
|  |    331 | done
 | 
|  |    332 | 
 | 
|  |    333 | lemma rtrancl_converseI: "<x,y>:converse(r^*) ==> <x,y>:converse(r)^*"
 | 
|  |    334 | apply (drule converseD)
 | 
|  |    335 | apply (frule rtrancl_type [THEN subsetD])
 | 
|  |    336 | apply (erule rtrancl_induct)
 | 
|  |    337 | apply (blast intro: rtrancl_refl)
 | 
|  |    338 | apply (blast intro: r_into_rtrancl rtrancl_trans)
 | 
|  |    339 | done
 | 
|  |    340 | 
 | 
|  |    341 | lemma rtrancl_converse: "converse(r)^* = converse(r^*)"
 | 
|  |    342 | apply (safe intro!: equalityI)
 | 
|  |    343 | apply (frule rtrancl_type [THEN subsetD])
 | 
|  |    344 | apply (safe dest!: rtrancl_converseD intro!: rtrancl_converseI)
 | 
|  |    345 | done
 | 
|  |    346 | 
 | 
|  |    347 | (** trancl **)
 | 
|  |    348 | 
 | 
|  |    349 | lemma trancl_converseD: "<a, b>:converse(r)^+ ==> <a, b>:converse(r^+)"
 | 
|  |    350 | apply (erule trancl_induct)
 | 
|  |    351 | apply (auto intro: r_into_trancl trancl_trans)
 | 
|  |    352 | done
 | 
|  |    353 | 
 | 
|  |    354 | lemma trancl_converseI: "<x,y>:converse(r^+) ==> <x,y>:converse(r)^+"
 | 
|  |    355 | apply (drule converseD)
 | 
|  |    356 | apply (erule trancl_induct)
 | 
|  |    357 | apply (auto intro: r_into_trancl trancl_trans)
 | 
|  |    358 | done
 | 
|  |    359 | 
 | 
|  |    360 | lemma trancl_converse: "converse(r)^+ = converse(r^+)"
 | 
|  |    361 | apply (safe intro!: equalityI)
 | 
|  |    362 | apply (frule trancl_type [THEN subsetD])
 | 
|  |    363 | apply (safe dest!: trancl_converseD intro!: trancl_converseI)
 | 
|  |    364 | done
 | 
|  |    365 | 
 | 
| 13534 |    366 | lemma converse_trancl_induct [case_names initial step, consumes 1]:
 | 
| 13239 |    367 | "[| <a, b>:r^+; !!y. <y, b> :r ==> P(y);  
 | 
|  |    368 |       !!y z. [| <y, z> : r; <z, b> : r^+; P(z) |] ==> P(y) |]  
 | 
|  |    369 |        ==> P(a)"
 | 
|  |    370 | apply (drule converseI)
 | 
|  |    371 | apply (simp (no_asm_use) add: trancl_converse [symmetric])
 | 
|  |    372 | apply (erule trancl_induct)
 | 
|  |    373 | apply (auto simp add: trancl_converse)
 | 
|  |    374 | done
 | 
|  |    375 | 
 | 
| 0 |    376 | end
 |