| author | wenzelm | 
| Thu, 24 Apr 1997 19:41:00 +0200 | |
| changeset 3046 | 6b7935317538 | 
| parent 3016 | 15763781afb0 | 
| child 4091 | 771b1f6422a8 | 
| 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  | 
||
11  | 
goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))";  | 
|
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  | 
|
17  | 
val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*";  | 
|
18  | 
by (rtac lfp_mono 1);  | 
|
19  | 
by (REPEAT (resolve_tac [rtrancl_bnd_mono, prem, 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: 
760 
diff
changeset
 | 
28  | 
bind_thm ("rtrancl_type", (rtrancl_def RS def_lfp_subset));
 | 
| 0 | 29  | 
|
30  | 
(*Reflexivity of rtrancl*)  | 
|
31  | 
val [prem] = goal Trancl.thy "[| a: field(r) |] ==> <a,a> : r^*";  | 
|
32  | 
by (resolve_tac [rtrancl_unfold RS ssubst] 1);  | 
|
33  | 
by (rtac (prem RS idI RS UnI1) 1);  | 
|
| 760 | 34  | 
qed "rtrancl_refl";  | 
| 0 | 35  | 
|
36  | 
(*Closure under composition with r *)  | 
|
37  | 
val prems = goal Trancl.thy  | 
|
38  | 
"[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^*";  | 
|
39  | 
by (resolve_tac [rtrancl_unfold RS ssubst] 1);  | 
|
40  | 
by (rtac (compI RS UnI2) 1);  | 
|
41  | 
by (resolve_tac prems 1);  | 
|
42  | 
by (resolve_tac prems 1);  | 
|
| 760 | 43  | 
qed "rtrancl_into_rtrancl";  | 
| 0 | 44  | 
|
45  | 
(*rtrancl of r contains all pairs in r *)  | 
|
46  | 
val prems = goal Trancl.thy "<a,b> : r ==> <a,b> : r^*";  | 
|
47  | 
by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1);  | 
|
48  | 
by (REPEAT (resolve_tac (prems@[fieldI1]) 1));  | 
|
| 760 | 49  | 
qed "r_into_rtrancl";  | 
| 0 | 50  | 
|
51  | 
(*The premise ensures that r consists entirely of pairs*)  | 
|
52  | 
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*";  | 
|
53  | 
by (cut_facts_tac prems 1);  | 
|
| 2929 | 54  | 
by (blast_tac (!claset addIs [r_into_rtrancl]) 1);  | 
| 760 | 55  | 
qed "r_subset_rtrancl";  | 
| 0 | 56  | 
|
57  | 
goal Trancl.thy "field(r^*) = field(r)";  | 
|
| 2929 | 58  | 
by (blast_tac (!claset addIs [r_into_rtrancl]  | 
| 1461 | 59  | 
addSDs [rtrancl_type RS subsetD]) 1);  | 
| 760 | 60  | 
qed "rtrancl_field";  | 
| 0 | 61  | 
|
62  | 
||
63  | 
(** standard induction rule **)  | 
|
64  | 
||
65  | 
val major::prems = goal Trancl.thy  | 
|
66  | 
"[| <a,b> : r^*; \  | 
|
67  | 
\ !!x. x: field(r) ==> P(<x,x>); \  | 
|
68  | 
\ !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |] \  | 
|
69  | 
\ ==> P(<a,b>)";  | 
|
70  | 
by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1);  | 
|
| 3016 | 71  | 
by (blast_tac (!claset addIs prems) 1);  | 
| 760 | 72  | 
qed "rtrancl_full_induct";  | 
| 0 | 73  | 
|
74  | 
(*nice induction rule.  | 
|
75  | 
Tried adding the typing hypotheses y,z:field(r), but these  | 
|
76  | 
caused expensive case splits!*)  | 
|
77  | 
val major::prems = goal Trancl.thy  | 
|
| 1461 | 78  | 
"[| <a,b> : r^*; \  | 
79  | 
\ P(a); \  | 
|
80  | 
\ !!y z.[| <a,y> : r^*; <y,z> : r; P(y) |] ==> P(z) \  | 
|
| 0 | 81  | 
\ |] ==> P(b)";  | 
82  | 
(*by induction on this formula*)  | 
|
83  | 
by (subgoal_tac "ALL y. <a,b> = <a,y> --> P(y)" 1);  | 
|
84  | 
(*now solve first subgoal: this formula is sufficient*)  | 
|
85  | 
by (EVERY1 [etac (spec RS mp), rtac refl]);  | 
|
86  | 
(*now do the induction*)  | 
|
87  | 
by (resolve_tac [major RS rtrancl_full_induct] 1);  | 
|
| 3016 | 88  | 
by (ALLGOALS (blast_tac (!claset addIs prems)));  | 
| 760 | 89  | 
qed "rtrancl_induct";  | 
| 0 | 90  | 
|
91  | 
(*transitivity of transitive closure!! -- by induction.*)  | 
|
92  | 
goalw Trancl.thy [trans_def] "trans(r^*)";  | 
|
93  | 
by (REPEAT (resolve_tac [allI,impI] 1));  | 
|
94  | 
by (eres_inst_tac [("b","z")] rtrancl_induct 1);
 | 
|
95  | 
by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1));  | 
|
| 760 | 96  | 
qed "trans_rtrancl";  | 
| 0 | 97  | 
|
98  | 
(*elimination of rtrancl -- by induction on a special formula*)  | 
|
99  | 
val major::prems = goal Trancl.thy  | 
|
| 1461 | 100  | 
"[| <a,b> : r^*; (a=b) ==> P; \  | 
101  | 
\ !!y.[| <a,y> : r^*; <y,b> : r |] ==> P |] \  | 
|
| 0 | 102  | 
\ ==> P";  | 
103  | 
by (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r)" 1);  | 
|
104  | 
(*see HOL/trancl*)  | 
|
105  | 
by (rtac (major RS rtrancl_induct) 2);  | 
|
| 2469 | 106  | 
by (ALLGOALS (fast_tac (!claset addSEs prems)));  | 
| 760 | 107  | 
qed "rtranclE";  | 
| 0 | 108  | 
|
109  | 
||
110  | 
(**** The relation trancl ****)  | 
|
111  | 
||
112  | 
(*Transitivity of r^+ is proved by transitivity of r^* *)  | 
|
113  | 
goalw Trancl.thy [trans_def,trancl_def] "trans(r^+)";  | 
|
| 3016 | 114  | 
by (blast_tac (!claset addIs [rtrancl_into_rtrancl RS  | 
115  | 
(trans_rtrancl RS transD RS compI)]) 1);  | 
|
| 760 | 116  | 
qed "trans_trancl";  | 
| 0 | 117  | 
|
118  | 
(** Conversions between trancl and rtrancl **)  | 
|
119  | 
||
| 3016 | 120  | 
goalw Trancl.thy [trancl_def] "!!r. <a,b> : r^+ ==> <a,b> : r^*";  | 
121  | 
by (blast_tac (!claset addIs [rtrancl_into_rtrancl]) 1);  | 
|
| 760 | 122  | 
qed "trancl_into_rtrancl";  | 
| 0 | 123  | 
|
124  | 
(*r^+ contains all pairs in r *)  | 
|
| 3016 | 125  | 
goalw Trancl.thy [trancl_def] "!!r. <a,b> : r ==> <a,b> : r^+";  | 
126  | 
by (blast_tac (!claset addSIs [rtrancl_refl]) 1);  | 
|
| 760 | 127  | 
qed "r_into_trancl";  | 
| 0 | 128  | 
|
129  | 
(*The premise ensures that r consists entirely of pairs*)  | 
|
| 3016 | 130  | 
goal Trancl.thy "!!r. r <= Sigma(A,B) ==> r <= r^+";  | 
| 2929 | 131  | 
by (blast_tac (!claset addIs [r_into_trancl]) 1);  | 
| 760 | 132  | 
qed "r_subset_trancl";  | 
| 0 | 133  | 
|
134  | 
(*intro rule by definition: from r^* and r *)  | 
|
| 3016 | 135  | 
goalw Trancl.thy [trancl_def]  | 
136  | 
"!!r. [| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+";  | 
|
137  | 
by (Blast_tac 1);  | 
|
| 760 | 138  | 
qed "rtrancl_into_trancl1";  | 
| 0 | 139  | 
|
140  | 
(*intro rule from r and r^* *)  | 
|
141  | 
val prems = goal Trancl.thy  | 
|
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);  | 
|
145  | 
by (etac (trans_trancl RS transD) 1);  | 
|
146  | 
by (etac r_into_trancl 1);  | 
|
| 760 | 147  | 
qed "rtrancl_into_trancl2";  | 
| 0 | 148  | 
|
149  | 
(*Nice induction rule for trancl*)  | 
|
150  | 
val major::prems = goal Trancl.thy  | 
|
| 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);  | 
| 2469 | 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*)  | 
|
165  | 
val major::prems = goal Trancl.thy  | 
|
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);  | 
|
| 2469 | 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);  | 
|
| 2929 | 174  | 
by (ALLGOALS (blast_tac (!claset addIs [rtrancl_into_trancl1])));  | 
| 760 | 175  | 
qed "tranclE";  | 
| 0 | 176  | 
|
177  | 
goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)";  | 
|
| 2929 | 178  | 
by (blast_tac (!claset addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);  | 
| 760 | 179  | 
qed "trancl_type";  | 
| 0 | 180  | 
|
181  | 
val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+";  | 
|
182  | 
by (REPEAT (resolve_tac [prem, comp_mono, rtrancl_mono] 1));  | 
|
| 760 | 183  | 
qed "trancl_mono";  | 
| 0 | 184  |