author | lcp |
Tue, 25 Apr 1995 11:14:03 +0200 | |
changeset 1072 | 0140ff702b23 |
parent 972 | e61b058d58d2 |
child 1121 | 485b49694253 |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/trancl |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
For trancl.thy. Theorems about the transitive closure of a relation |
|
7 |
*) |
|
8 |
||
9 |
open Trancl; |
|
10 |
||
11 |
(** Natural deduction for trans(r) **) |
|
12 |
||
13 |
val prems = goalw Trancl.thy [trans_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
14 |
"(!! x y z. [| (x,y):r; (y,z):r |] ==> (x,z):r) ==> trans(r)"; |
923 | 15 |
by (REPEAT (ares_tac (prems@[allI,impI]) 1)); |
16 |
qed "transI"; |
|
17 |
||
18 |
val major::prems = goalw Trancl.thy [trans_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
19 |
"[| trans(r); (a,b):r; (b,c):r |] ==> (a,c):r"; |
923 | 20 |
by (cut_facts_tac [major] 1); |
21 |
by (fast_tac (HOL_cs addIs prems) 1); |
|
22 |
qed "transD"; |
|
23 |
||
24 |
(** Identity relation **) |
|
25 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
26 |
goalw Trancl.thy [id_def] "(a,a) : id"; |
923 | 27 |
by (rtac CollectI 1); |
28 |
by (rtac exI 1); |
|
29 |
by (rtac refl 1); |
|
30 |
qed "idI"; |
|
31 |
||
32 |
val major::prems = goalw Trancl.thy [id_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
33 |
"[| p: id; !!x.[| p = (x,x) |] ==> P \ |
923 | 34 |
\ |] ==> P"; |
35 |
by (rtac (major RS CollectE) 1); |
|
36 |
by (etac exE 1); |
|
37 |
by (eresolve_tac prems 1); |
|
38 |
qed "idE"; |
|
39 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
40 |
goalw Trancl.thy [id_def] "(a,b):id = (a=b)"; |
923 | 41 |
by(fast_tac prod_cs 1); |
42 |
qed "pair_in_id_conv"; |
|
43 |
||
44 |
(** Composition of two relations **) |
|
45 |
||
46 |
val prems = goalw Trancl.thy [comp_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
47 |
"[| (a,b):s; (b,c):r |] ==> (a,c) : r O s"; |
923 | 48 |
by (fast_tac (set_cs addIs prems) 1); |
49 |
qed "compI"; |
|
50 |
||
51 |
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) |
|
52 |
val prems = goalw Trancl.thy [comp_def] |
|
53 |
"[| xz : r O s; \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
54 |
\ !!x y z. [| xz = (x,z); (x,y):s; (y,z):r |] ==> P \ |
923 | 55 |
\ |] ==> P"; |
56 |
by (cut_facts_tac prems 1); |
|
57 |
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); |
|
58 |
qed "compE"; |
|
59 |
||
60 |
val prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
61 |
"[| (a,c) : r O s; \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
62 |
\ !!y. [| (a,y):s; (y,c):r |] ==> P \ |
923 | 63 |
\ |] ==> P"; |
64 |
by (rtac compE 1); |
|
65 |
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1)); |
|
66 |
qed "compEpair"; |
|
67 |
||
68 |
val comp_cs = prod_cs addIs [compI, idI] addSEs [compE, idE]; |
|
69 |
||
70 |
goal Trancl.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; |
|
71 |
by (fast_tac comp_cs 1); |
|
72 |
qed "comp_mono"; |
|
73 |
||
74 |
goal Trancl.thy |
|
75 |
"!!r s. [| s <= Sigma A (%x.B); r <= Sigma B (%x.C) |] ==> \ |
|
76 |
\ (r O s) <= Sigma A (%x.C)"; |
|
77 |
by (fast_tac comp_cs 1); |
|
78 |
qed "comp_subset_Sigma"; |
|
79 |
||
80 |
||
81 |
(** The relation rtrancl **) |
|
82 |
||
83 |
goal Trancl.thy "mono(%s. id Un (r O s))"; |
|
84 |
by (rtac monoI 1); |
|
85 |
by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1)); |
|
86 |
qed "rtrancl_fun_mono"; |
|
87 |
||
88 |
val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); |
|
89 |
||
90 |
(*Reflexivity of rtrancl*) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
91 |
goal Trancl.thy "(a,a) : r^*"; |
923 | 92 |
by (stac rtrancl_unfold 1); |
93 |
by (fast_tac comp_cs 1); |
|
94 |
qed "rtrancl_refl"; |
|
95 |
||
96 |
(*Closure under composition with r*) |
|
97 |
val prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
98 |
"[| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^*"; |
923 | 99 |
by (stac rtrancl_unfold 1); |
100 |
by (fast_tac (comp_cs addIs prems) 1); |
|
101 |
qed "rtrancl_into_rtrancl"; |
|
102 |
||
103 |
(*rtrancl of r contains r*) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
104 |
val [prem] = goal Trancl.thy "[| (a,b) : r |] ==> (a,b) : r^*"; |
923 | 105 |
by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1); |
106 |
by (rtac prem 1); |
|
107 |
qed "r_into_rtrancl"; |
|
108 |
||
109 |
(*monotonicity of rtrancl*) |
|
110 |
goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*"; |
|
111 |
by(REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1)); |
|
112 |
qed "rtrancl_mono"; |
|
113 |
||
114 |
(** standard induction rule **) |
|
115 |
||
116 |
val major::prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
117 |
"[| (a,b) : r^*; \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
118 |
\ !!x. P((x,x)); \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
119 |
\ !!x y z.[| P((x,y)); (x,y): r^*; (y,z): r |] ==> P((x,z)) |] \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
120 |
\ ==> P((a,b))"; |
923 | 121 |
by (rtac ([rtrancl_def, rtrancl_fun_mono, major] MRS def_induct) 1); |
122 |
by (fast_tac (comp_cs addIs prems) 1); |
|
123 |
qed "rtrancl_full_induct"; |
|
124 |
||
125 |
(*nice induction rule*) |
|
126 |
val major::prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
127 |
"[| (a::'a,b) : r^*; \ |
923 | 128 |
\ P(a); \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
129 |
\ !!y z.[| (a,y) : r^*; (y,z) : r; P(y) |] ==> P(z) |] \ |
923 | 130 |
\ ==> P(b)"; |
131 |
(*by induction on this formula*) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
132 |
by (subgoal_tac "! y. (a::'a,b) = (a,y) --> P(y)" 1); |
923 | 133 |
(*now solve first subgoal: this formula is sufficient*) |
134 |
by (fast_tac HOL_cs 1); |
|
135 |
(*now do the induction*) |
|
136 |
by (resolve_tac [major RS rtrancl_full_induct] 1); |
|
137 |
by (fast_tac (comp_cs addIs prems) 1); |
|
138 |
by (fast_tac (comp_cs addIs prems) 1); |
|
139 |
qed "rtrancl_induct"; |
|
140 |
||
141 |
(*transitivity of transitive closure!! -- by induction.*) |
|
142 |
goal Trancl.thy "trans(r^*)"; |
|
143 |
by (rtac transI 1); |
|
144 |
by (res_inst_tac [("b","z")] rtrancl_induct 1); |
|
145 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); |
|
146 |
qed "trans_rtrancl"; |
|
147 |
||
148 |
(*elimination of rtrancl -- by induction on a special formula*) |
|
149 |
val major::prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
150 |
"[| (a::'a,b) : r^*; (a = b) ==> P; \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
151 |
\ !!y.[| (a,y) : r^*; (y,b) : r |] ==> P \ |
923 | 152 |
\ |] ==> P"; |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
153 |
by (subgoal_tac "(a::'a) = b | (? y. (a,y) : r^* & (y,b) : r)" 1); |
923 | 154 |
by (rtac (major RS rtrancl_induct) 2); |
155 |
by (fast_tac (set_cs addIs prems) 2); |
|
156 |
by (fast_tac (set_cs addIs prems) 2); |
|
157 |
by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); |
|
158 |
qed "rtranclE"; |
|
159 |
||
160 |
||
161 |
(**** The relation trancl ****) |
|
162 |
||
163 |
(** Conversions between trancl and rtrancl **) |
|
164 |
||
165 |
val [major] = goalw Trancl.thy [trancl_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
166 |
"(a,b) : r^+ ==> (a,b) : r^*"; |
923 | 167 |
by (resolve_tac [major RS compEpair] 1); |
168 |
by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); |
|
169 |
qed "trancl_into_rtrancl"; |
|
170 |
||
171 |
(*r^+ contains r*) |
|
172 |
val [prem] = goalw Trancl.thy [trancl_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
173 |
"[| (a,b) : r |] ==> (a,b) : r^+"; |
923 | 174 |
by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); |
175 |
qed "r_into_trancl"; |
|
176 |
||
177 |
(*intro rule by definition: from rtrancl and r*) |
|
178 |
val prems = goalw Trancl.thy [trancl_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
179 |
"[| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^+"; |
923 | 180 |
by (REPEAT (resolve_tac ([compI]@prems) 1)); |
181 |
qed "rtrancl_into_trancl1"; |
|
182 |
||
183 |
(*intro rule from r and rtrancl*) |
|
184 |
val prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
185 |
"[| (a,b) : r; (b,c) : r^* |] ==> (a,c) : r^+"; |
923 | 186 |
by (resolve_tac (prems RL [rtranclE]) 1); |
187 |
by (etac subst 1); |
|
188 |
by (resolve_tac (prems RL [r_into_trancl]) 1); |
|
189 |
by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1); |
|
190 |
by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1)); |
|
191 |
qed "rtrancl_into_trancl2"; |
|
192 |
||
193 |
(*elimination of r^+ -- NOT an induction rule*) |
|
194 |
val major::prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
195 |
"[| (a::'a,b) : r^+; \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
196 |
\ (a,b) : r ==> P; \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
197 |
\ !!y.[| (a,y) : r^+; (y,b) : r |] ==> P \ |
923 | 198 |
\ |] ==> P"; |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
199 |
by (subgoal_tac "(a::'a,b) : r | (? y. (a,y) : r^+ & (y,b) : r)" 1); |
923 | 200 |
by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1)); |
201 |
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); |
|
202 |
by (etac rtranclE 1); |
|
203 |
by (fast_tac comp_cs 1); |
|
204 |
by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1); |
|
205 |
qed "tranclE"; |
|
206 |
||
207 |
(*Transitivity of r^+. |
|
208 |
Proved by unfolding since it uses transitivity of rtrancl. *) |
|
209 |
goalw Trancl.thy [trancl_def] "trans(r^+)"; |
|
210 |
by (rtac transI 1); |
|
211 |
by (REPEAT (etac compEpair 1)); |
|
212 |
by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); |
|
213 |
by (REPEAT (assume_tac 1)); |
|
214 |
qed "trans_trancl"; |
|
215 |
||
216 |
val prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
217 |
"[| (a,b) : r; (b,c) : r^+ |] ==> (a,c) : r^+"; |
923 | 218 |
by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1); |
219 |
by (resolve_tac prems 1); |
|
220 |
by (resolve_tac prems 1); |
|
221 |
qed "trancl_into_trancl2"; |
|
222 |
||
223 |
||
224 |
val major::prems = goal Trancl.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
225 |
"[| (a,b) : r^*; r <= Sigma A (%x.A) |] ==> a=b | a:A"; |
923 | 226 |
by (cut_facts_tac prems 1); |
227 |
by (rtac (major RS rtrancl_induct) 1); |
|
228 |
by (rtac (refl RS disjI1) 1); |
|
229 |
by (fast_tac (comp_cs addSEs [SigmaE2]) 1); |
|
230 |
qed "trancl_subset_Sigma_lemma"; |
|
231 |
||
232 |
goalw Trancl.thy [trancl_def] |
|
233 |
"!!r. r <= Sigma A (%x.A) ==> trancl(r) <= Sigma A (%x.A)"; |
|
234 |
by (fast_tac (comp_cs addSDs [trancl_subset_Sigma_lemma]) 1); |
|
235 |
qed "trancl_subset_Sigma"; |
|
236 |
||
237 |
val prod_ss = prod_ss addsimps [pair_in_id_conv]; |