author | nipkow |
Fri, 28 Jul 1995 18:05:50 +0200 | |
changeset 1212 | 7059356e18e0 |
parent 782 | 200a16083201 |
child 1461 | 6bcb44e4d6e5 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: ZF/trancl.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
For trancl.thy. Transitive closure of a relation |
|
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)); |
|
14 |
by (fast_tac comp_cs 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, |
|
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); |
|
54 |
by (fast_tac (ZF_cs addIs [r_into_rtrancl]) 1); |
|
760 | 55 |
qed "r_subset_rtrancl"; |
0 | 56 |
|
57 |
goal Trancl.thy "field(r^*) = field(r)"; |
|
58 |
by (fast_tac (eq_cs addIs [r_into_rtrancl] |
|
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); |
|
71 |
by (fast_tac (ZF_cs addIs prems addSEs [idE,compE]) 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 |
|
78 |
"[| <a,b> : r^*; \ |
|
79 |
\ P(a); \ |
|
80 |
\ !!y z.[| <a,y> : r^*; <y,z> : r; P(y) |] ==> P(z) \ |
|
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); |
|
88 |
by (ALLGOALS (fast_tac (ZF_cs 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 |
|
100 |
"[| <a,b> : r^*; (a=b) ==> P; \ |
|
101 |
\ !!y.[| <a,y> : r^*; <y,b> : r |] ==> P |] \ |
|
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); |
|
106 |
by (ALLGOALS (fast_tac (ZF_cs 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^+)"; |
|
114 |
by (safe_tac comp_cs); |
|
115 |
by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); |
|
116 |
by (REPEAT (assume_tac 1)); |
|
760 | 117 |
qed "trans_trancl"; |
0 | 118 |
|
119 |
(** Conversions between trancl and rtrancl **) |
|
120 |
||
121 |
val [major] = goalw Trancl.thy [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*"; |
|
122 |
by (resolve_tac [major RS compEpair] 1); |
|
123 |
by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); |
|
760 | 124 |
qed "trancl_into_rtrancl"; |
0 | 125 |
|
126 |
(*r^+ contains all pairs in r *) |
|
127 |
val [prem] = goalw Trancl.thy [trancl_def] "<a,b> : r ==> <a,b> : r^+"; |
|
128 |
by (REPEAT (ares_tac [prem,compI,rtrancl_refl,fieldI1] 1)); |
|
760 | 129 |
qed "r_into_trancl"; |
0 | 130 |
|
131 |
(*The premise ensures that r consists entirely of pairs*) |
|
132 |
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+"; |
|
133 |
by (cut_facts_tac prems 1); |
|
134 |
by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); |
|
760 | 135 |
qed "r_subset_trancl"; |
0 | 136 |
|
137 |
(*intro rule by definition: from r^* and r *) |
|
138 |
val prems = goalw Trancl.thy [trancl_def] |
|
139 |
"[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+"; |
|
140 |
by (REPEAT (resolve_tac ([compI]@prems) 1)); |
|
760 | 141 |
qed "rtrancl_into_trancl1"; |
0 | 142 |
|
143 |
(*intro rule from r and r^* *) |
|
144 |
val prems = goal Trancl.thy |
|
145 |
"[| <a,b> : r; <b,c> : r^* |] ==> <a,c> : r^+"; |
|
146 |
by (resolve_tac (prems RL [rtrancl_induct]) 1); |
|
147 |
by (resolve_tac (prems RL [r_into_trancl]) 1); |
|
148 |
by (etac (trans_trancl RS transD) 1); |
|
149 |
by (etac r_into_trancl 1); |
|
760 | 150 |
qed "rtrancl_into_trancl2"; |
0 | 151 |
|
152 |
(*Nice induction rule for trancl*) |
|
153 |
val major::prems = goal Trancl.thy |
|
154 |
"[| <a,b> : r^+; \ |
|
155 |
\ !!y. [| <a,y> : r |] ==> P(y); \ |
|
156 |
\ !!y z.[| <a,y> : r^+; <y,z> : r; P(y) |] ==> P(z) \ |
|
157 |
\ |] ==> P(b)"; |
|
158 |
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); |
|
159 |
(*by induction on this formula*) |
|
160 |
by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1); |
|
161 |
(*now solve first subgoal: this formula is sufficient*) |
|
162 |
by (fast_tac ZF_cs 1); |
|
163 |
by (etac rtrancl_induct 1); |
|
164 |
by (ALLGOALS (fast_tac (ZF_cs addIs (rtrancl_into_trancl1::prems)))); |
|
760 | 165 |
qed "trancl_induct"; |
0 | 166 |
|
167 |
(*elimination of r^+ -- NOT an induction rule*) |
|
168 |
val major::prems = goal Trancl.thy |
|
169 |
"[| <a,b> : r^+; \ |
|
170 |
\ <a,b> : r ==> P; \ |
|
171 |
\ !!y.[| <a,y> : r^+; <y,b> : r |] ==> P \ |
|
172 |
\ |] ==> P"; |
|
173 |
by (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r)" 1); |
|
174 |
by (fast_tac (ZF_cs addIs prems) 1); |
|
175 |
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); |
|
176 |
by (etac rtranclE 1); |
|
177 |
by (ALLGOALS (fast_tac (ZF_cs addIs [rtrancl_into_trancl1]))); |
|
760 | 178 |
qed "tranclE"; |
0 | 179 |
|
180 |
goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)"; |
|
181 |
by (fast_tac (ZF_cs addEs [compE, rtrancl_type RS subsetD RS SigmaE2]) 1); |
|
760 | 182 |
qed "trancl_type"; |
0 | 183 |
|
184 |
val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+"; |
|
185 |
by (REPEAT (resolve_tac [prem, comp_mono, rtrancl_mono] 1)); |
|
760 | 186 |
qed "trancl_mono"; |
0 | 187 |