| author | paulson |
| Thu, 30 Jan 2003 10:35:56 +0100 | |
| changeset 13796 | 19f50fa807ae |
| parent 9685 | 6d123a7e30bd |
| child 13798 | 4c1a53627500 |
| permissions | -rw-r--r-- |
| 4776 | 1 |
(* Title: HOL/UNITY/SubstAx |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
| 6536 | 6 |
Weak LeadsTo relation (restricted to the set of reachable states) |
| 4776 | 7 |
*) |
8 |
||
| 13796 | 9 |
theory SubstAx = WFair + Constrains: |
| 4776 | 10 |
|
| 8041 | 11 |
constdefs |
| 13796 | 12 |
Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) |
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8041
diff
changeset
|
13 |
"A Ensures B == {F. F : (reachable F Int A) ensures B}"
|
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8041
diff
changeset
|
14 |
|
| 13796 | 15 |
LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) |
| 8041 | 16 |
"A LeadsTo B == {F. F : (reachable F Int A) leadsTo B}"
|
| 4776 | 17 |
|
| 9685 | 18 |
syntax (xsymbols) |
| 13796 | 19 |
"op LeadsTo" :: "['a set, 'a set] => 'a program set" (infixl " \<longmapsto>w " 60) |
20 |
||
21 |
||
22 |
(*Resembles the previous definition of LeadsTo*) |
|
23 |
lemma LeadsTo_eq_leadsTo: |
|
24 |
"A LeadsTo B = {F. F : (reachable F Int A) leadsTo (reachable F Int B)}"
|
|
25 |
apply (unfold LeadsTo_def) |
|
26 |
apply (blast dest: psp_stable2 intro: leadsTo_weaken) |
|
27 |
done |
|
28 |
||
29 |
||
30 |
(*** Specialized laws for handling invariants ***) |
|
31 |
||
32 |
(** Conjoining an Always property **) |
|
33 |
||
34 |
lemma Always_LeadsTo_pre: |
|
35 |
"F : Always INV ==> (F : (INV Int A) LeadsTo A') = (F : A LeadsTo A')" |
|
36 |
by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric]) |
|
37 |
||
38 |
lemma Always_LeadsTo_post: |
|
39 |
"F : Always INV ==> (F : A LeadsTo (INV Int A')) = (F : A LeadsTo A')" |
|
40 |
by (simp add: LeadsTo_eq_leadsTo Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric]) |
|
41 |
||
42 |
(* [| F : Always C; F : (C Int A) LeadsTo A' |] ==> F : A LeadsTo A' *) |
|
43 |
lemmas Always_LeadsToI = Always_LeadsTo_pre [THEN iffD1, standard] |
|
44 |
||
45 |
(* [| F : Always INV; F : A LeadsTo A' |] ==> F : A LeadsTo (INV Int A') *) |
|
46 |
lemmas Always_LeadsToD = Always_LeadsTo_post [THEN iffD2, standard] |
|
47 |
||
48 |
||
49 |
(*** Introduction rules: Basis, Trans, Union ***) |
|
50 |
||
51 |
lemma leadsTo_imp_LeadsTo: "F : A leadsTo B ==> F : A LeadsTo B" |
|
52 |
apply (simp add: LeadsTo_def) |
|
53 |
apply (blast intro: leadsTo_weaken_L) |
|
54 |
done |
|
55 |
||
56 |
lemma LeadsTo_Trans: |
|
57 |
"[| F : A LeadsTo B; F : B LeadsTo C |] ==> F : A LeadsTo C" |
|
58 |
apply (simp add: LeadsTo_eq_leadsTo) |
|
59 |
apply (blast intro: leadsTo_Trans) |
|
60 |
done |
|
61 |
||
62 |
lemma LeadsTo_Union: |
|
63 |
"(!!A. A : S ==> F : A LeadsTo B) ==> F : (Union S) LeadsTo B" |
|
64 |
apply (simp add: LeadsTo_def) |
|
65 |
apply (subst Int_Union) |
|
66 |
apply (blast intro: leadsTo_UN) |
|
67 |
done |
|
68 |
||
69 |
||
70 |
(*** Derived rules ***) |
|
71 |
||
72 |
lemma LeadsTo_UNIV [simp]: "F : A LeadsTo UNIV" |
|
73 |
by (simp add: LeadsTo_def) |
|
74 |
||
75 |
(*Useful with cancellation, disjunction*) |
|
76 |
lemma LeadsTo_Un_duplicate: |
|
77 |
"F : A LeadsTo (A' Un A') ==> F : A LeadsTo A'" |
|
78 |
by (simp add: Un_ac) |
|
79 |
||
80 |
lemma LeadsTo_Un_duplicate2: |
|
81 |
"F : A LeadsTo (A' Un C Un C) ==> F : A LeadsTo (A' Un C)" |
|
82 |
by (simp add: Un_ac) |
|
83 |
||
84 |
lemma LeadsTo_UN: |
|
85 |
"(!!i. i : I ==> F : (A i) LeadsTo B) ==> F : (UN i:I. A i) LeadsTo B" |
|
86 |
apply (simp only: Union_image_eq [symmetric]) |
|
87 |
apply (blast intro: LeadsTo_Union) |
|
88 |
done |
|
89 |
||
90 |
(*Binary union introduction rule*) |
|
91 |
lemma LeadsTo_Un: |
|
92 |
"[| F : A LeadsTo C; F : B LeadsTo C |] ==> F : (A Un B) LeadsTo C" |
|
93 |
apply (subst Un_eq_Union) |
|
94 |
apply (blast intro: LeadsTo_Union) |
|
95 |
done |
|
96 |
||
97 |
(*Lets us look at the starting state*) |
|
98 |
lemma single_LeadsTo_I: |
|
99 |
"(!!s. s : A ==> F : {s} LeadsTo B) ==> F : A LeadsTo B"
|
|
100 |
by (subst UN_singleton [symmetric], rule LeadsTo_UN, blast) |
|
101 |
||
102 |
lemma subset_imp_LeadsTo: "A <= B ==> F : A LeadsTo B" |
|
103 |
apply (simp add: LeadsTo_def) |
|
104 |
apply (blast intro: subset_imp_leadsTo) |
|
105 |
done |
|
106 |
||
107 |
lemmas empty_LeadsTo = empty_subsetI [THEN subset_imp_LeadsTo, standard, simp] |
|
108 |
||
109 |
lemma LeadsTo_weaken_R [rule_format]: |
|
110 |
"[| F : A LeadsTo A'; A' <= B' |] ==> F : A LeadsTo B'" |
|
111 |
apply (simp (no_asm_use) add: LeadsTo_def) |
|
112 |
apply (blast intro: leadsTo_weaken_R) |
|
113 |
done |
|
114 |
||
115 |
lemma LeadsTo_weaken_L [rule_format]: |
|
116 |
"[| F : A LeadsTo A'; B <= A |] |
|
117 |
==> F : B LeadsTo A'" |
|
118 |
apply (simp (no_asm_use) add: LeadsTo_def) |
|
119 |
apply (blast intro: leadsTo_weaken_L) |
|
120 |
done |
|
121 |
||
122 |
lemma LeadsTo_weaken: |
|
123 |
"[| F : A LeadsTo A'; |
|
124 |
B <= A; A' <= B' |] |
|
125 |
==> F : B LeadsTo B'" |
|
126 |
by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans) |
|
127 |
||
128 |
lemma Always_LeadsTo_weaken: |
|
129 |
"[| F : Always C; F : A LeadsTo A'; |
|
130 |
C Int B <= A; C Int A' <= B' |] |
|
131 |
==> F : B LeadsTo B'" |
|
132 |
by (blast dest: Always_LeadsToI intro: LeadsTo_weaken intro: Always_LeadsToD) |
|
133 |
||
134 |
(** Two theorems for "proof lattices" **) |
|
135 |
||
136 |
lemma LeadsTo_Un_post: "F : A LeadsTo B ==> F : (A Un B) LeadsTo B" |
|
137 |
by (blast intro: LeadsTo_Un subset_imp_LeadsTo) |
|
138 |
||
139 |
lemma LeadsTo_Trans_Un: |
|
140 |
"[| F : A LeadsTo B; F : B LeadsTo C |] |
|
141 |
==> F : (A Un B) LeadsTo C" |
|
142 |
by (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans) |
|
143 |
||
144 |
||
145 |
(** Distributive laws **) |
|
146 |
||
147 |
lemma LeadsTo_Un_distrib: |
|
148 |
"(F : (A Un B) LeadsTo C) = (F : A LeadsTo C & F : B LeadsTo C)" |
|
149 |
by (blast intro: LeadsTo_Un LeadsTo_weaken_L) |
|
150 |
||
151 |
lemma LeadsTo_UN_distrib: |
|
152 |
"(F : (UN i:I. A i) LeadsTo B) = (ALL i : I. F : (A i) LeadsTo B)" |
|
153 |
by (blast intro: LeadsTo_UN LeadsTo_weaken_L) |
|
154 |
||
155 |
lemma LeadsTo_Union_distrib: |
|
156 |
"(F : (Union S) LeadsTo B) = (ALL A : S. F : A LeadsTo B)" |
|
157 |
by (blast intro: LeadsTo_Union LeadsTo_weaken_L) |
|
158 |
||
159 |
||
160 |
(** More rules using the premise "Always INV" **) |
|
161 |
||
162 |
lemma LeadsTo_Basis: "F : A Ensures B ==> F : A LeadsTo B" |
|
163 |
by (simp add: Ensures_def LeadsTo_def leadsTo_Basis) |
|
164 |
||
165 |
lemma EnsuresI: |
|
166 |
"[| F : (A-B) Co (A Un B); F : transient (A-B) |] |
|
167 |
==> F : A Ensures B" |
|
168 |
apply (simp add: Ensures_def Constrains_eq_constrains) |
|
169 |
apply (blast intro: ensuresI constrains_weaken transient_strengthen) |
|
170 |
done |
|
171 |
||
172 |
lemma Always_LeadsTo_Basis: |
|
173 |
"[| F : Always INV; |
|
174 |
F : (INV Int (A-A')) Co (A Un A'); |
|
175 |
F : transient (INV Int (A-A')) |] |
|
176 |
==> F : A LeadsTo A'" |
|
177 |
apply (rule Always_LeadsToI, assumption) |
|
178 |
apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) |
|
179 |
done |
|
180 |
||
181 |
(*Set difference: maybe combine with leadsTo_weaken_L?? |
|
182 |
This is the most useful form of the "disjunction" rule*) |
|
183 |
lemma LeadsTo_Diff: |
|
184 |
"[| F : (A-B) LeadsTo C; F : (A Int B) LeadsTo C |] |
|
185 |
==> F : A LeadsTo C" |
|
186 |
by (blast intro: LeadsTo_Un LeadsTo_weaken) |
|
187 |
||
188 |
||
189 |
lemma LeadsTo_UN_UN: |
|
190 |
"(!! i. i:I ==> F : (A i) LeadsTo (A' i)) |
|
191 |
==> F : (UN i:I. A i) LeadsTo (UN i:I. A' i)" |
|
192 |
apply (simp only: Union_image_eq [symmetric]) |
|
193 |
apply (blast intro: LeadsTo_Union LeadsTo_weaken_R) |
|
194 |
done |
|
195 |
||
196 |
||
197 |
(*Version with no index set*) |
|
198 |
lemma LeadsTo_UN_UN_noindex: |
|
199 |
"(!! i. F : (A i) LeadsTo (A' i)) |
|
200 |
==> F : (UN i. A i) LeadsTo (UN i. A' i)" |
|
201 |
by (blast intro: LeadsTo_UN_UN) |
|
202 |
||
203 |
(*Version with no index set*) |
|
204 |
lemma all_LeadsTo_UN_UN: |
|
205 |
"ALL i. F : (A i) LeadsTo (A' i) |
|
206 |
==> F : (UN i. A i) LeadsTo (UN i. A' i)" |
|
207 |
by (blast intro: LeadsTo_UN_UN) |
|
208 |
||
209 |
(*Binary union version*) |
|
210 |
lemma LeadsTo_Un_Un: |
|
211 |
"[| F : A LeadsTo A'; F : B LeadsTo B' |] |
|
212 |
==> F : (A Un B) LeadsTo (A' Un B')" |
|
213 |
by (blast intro: LeadsTo_Un LeadsTo_weaken_R) |
|
214 |
||
215 |
||
216 |
(** The cancellation law **) |
|
217 |
||
218 |
lemma LeadsTo_cancel2: |
|
219 |
"[| F : A LeadsTo (A' Un B); F : B LeadsTo B' |] |
|
220 |
==> F : A LeadsTo (A' Un B')" |
|
221 |
by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans) |
|
222 |
||
223 |
lemma LeadsTo_cancel_Diff2: |
|
224 |
"[| F : A LeadsTo (A' Un B); F : (B-A') LeadsTo B' |] |
|
225 |
==> F : A LeadsTo (A' Un B')" |
|
226 |
apply (rule LeadsTo_cancel2) |
|
227 |
prefer 2 apply assumption |
|
228 |
apply (simp_all (no_asm_simp)) |
|
229 |
done |
|
230 |
||
231 |
lemma LeadsTo_cancel1: |
|
232 |
"[| F : A LeadsTo (B Un A'); F : B LeadsTo B' |] |
|
233 |
==> F : A LeadsTo (B' Un A')" |
|
234 |
apply (simp add: Un_commute) |
|
235 |
apply (blast intro!: LeadsTo_cancel2) |
|
236 |
done |
|
237 |
||
238 |
lemma LeadsTo_cancel_Diff1: |
|
239 |
"[| F : A LeadsTo (B Un A'); F : (B-A') LeadsTo B' |] |
|
240 |
==> F : A LeadsTo (B' Un A')" |
|
241 |
apply (rule LeadsTo_cancel1) |
|
242 |
prefer 2 apply assumption |
|
243 |
apply (simp_all (no_asm_simp)) |
|
244 |
done |
|
245 |
||
246 |
||
247 |
(** The impossibility law **) |
|
248 |
||
249 |
(*The set "A" may be non-empty, but it contains no reachable states*) |
|
250 |
lemma LeadsTo_empty: "F : A LeadsTo {} ==> F : Always (-A)"
|
|
251 |
apply (simp (no_asm_use) add: LeadsTo_def Always_eq_includes_reachable) |
|
252 |
apply (drule leadsTo_empty, auto) |
|
253 |
done |
|
254 |
||
255 |
||
256 |
(** PSP: Progress-Safety-Progress **) |
|
257 |
||
258 |
(*Special case of PSP: Misra's "stable conjunction"*) |
|
259 |
lemma PSP_Stable: |
|
260 |
"[| F : A LeadsTo A'; F : Stable B |] |
|
261 |
==> F : (A Int B) LeadsTo (A' Int B)" |
|
262 |
apply (simp (no_asm_use) add: LeadsTo_eq_leadsTo Stable_eq_stable) |
|
263 |
apply (drule psp_stable, assumption) |
|
264 |
apply (simp add: Int_ac) |
|
265 |
done |
|
266 |
||
267 |
lemma PSP_Stable2: |
|
268 |
"[| F : A LeadsTo A'; F : Stable B |] |
|
269 |
==> F : (B Int A) LeadsTo (B Int A')" |
|
270 |
by (simp add: PSP_Stable Int_ac) |
|
271 |
||
272 |
lemma PSP: |
|
273 |
"[| F : A LeadsTo A'; F : B Co B' |] |
|
274 |
==> F : (A Int B') LeadsTo ((A' Int B) Un (B' - B))" |
|
275 |
apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains) |
|
276 |
apply (blast dest: psp intro: leadsTo_weaken) |
|
277 |
done |
|
278 |
||
279 |
lemma PSP2: |
|
280 |
"[| F : A LeadsTo A'; F : B Co B' |] |
|
281 |
==> F : (B' Int A) LeadsTo ((B Int A') Un (B' - B))" |
|
282 |
by (simp add: PSP Int_ac) |
|
283 |
||
284 |
lemma PSP_Unless: |
|
285 |
"[| F : A LeadsTo A'; F : B Unless B' |] |
|
286 |
==> F : (A Int B) LeadsTo ((A' Int B) Un B')" |
|
287 |
apply (unfold Unless_def) |
|
288 |
apply (drule PSP, assumption) |
|
289 |
apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo) |
|
290 |
done |
|
291 |
||
292 |
||
293 |
lemma Stable_transient_Always_LeadsTo: |
|
294 |
"[| F : Stable A; F : transient C; |
|
295 |
F : Always (-A Un B Un C) |] ==> F : A LeadsTo B" |
|
296 |
apply (erule Always_LeadsTo_weaken) |
|
297 |
apply (rule LeadsTo_Diff) |
|
298 |
prefer 2 |
|
299 |
apply (erule |
|
300 |
transient_imp_leadsTo [THEN leadsTo_imp_LeadsTo, THEN PSP_Stable2]) |
|
301 |
apply (blast intro: subset_imp_LeadsTo)+ |
|
302 |
done |
|
303 |
||
304 |
||
305 |
(*** Induction rules ***) |
|
306 |
||
307 |
(** Meta or object quantifier ????? **) |
|
308 |
lemma LeadsTo_wf_induct: |
|
309 |
"[| wf r; |
|
310 |
ALL m. F : (A Int f-`{m}) LeadsTo
|
|
311 |
((A Int f-`(r^-1 `` {m})) Un B) |]
|
|
312 |
==> F : A LeadsTo B" |
|
313 |
apply (simp (no_asm_use) add: LeadsTo_eq_leadsTo) |
|
314 |
apply (erule leadsTo_wf_induct) |
|
315 |
apply (blast intro: leadsTo_weaken) |
|
316 |
done |
|
317 |
||
318 |
||
319 |
lemma Bounded_induct: |
|
320 |
"[| wf r; |
|
321 |
ALL m:I. F : (A Int f-`{m}) LeadsTo
|
|
322 |
((A Int f-`(r^-1 `` {m})) Un B) |]
|
|
323 |
==> F : A LeadsTo ((A - (f-`I)) Un B)" |
|
324 |
apply (erule LeadsTo_wf_induct, safe) |
|
325 |
apply (case_tac "m:I") |
|
326 |
apply (blast intro: LeadsTo_weaken) |
|
327 |
apply (blast intro: subset_imp_LeadsTo) |
|
328 |
done |
|
329 |
||
330 |
||
331 |
lemma LessThan_induct: |
|
332 |
"(!!m::nat. F : (A Int f-`{m}) LeadsTo ((A Int f-`(lessThan m)) Un B))
|
|
333 |
==> F : A LeadsTo B" |
|
334 |
apply (rule wf_less_than [THEN LeadsTo_wf_induct], auto) |
|
335 |
done |
|
336 |
||
337 |
(*Integer version. Could generalize from 0 to any lower bound*) |
|
338 |
lemma integ_0_le_induct: |
|
339 |
"[| F : Always {s. (0::int) <= f s};
|
|
340 |
!! z. F : (A Int {s. f s = z}) LeadsTo
|
|
341 |
((A Int {s. f s < z}) Un B) |]
|
|
342 |
==> F : A LeadsTo B" |
|
343 |
apply (rule_tac f = "nat o f" in LessThan_induct) |
|
344 |
apply (simp add: vimage_def) |
|
345 |
apply (rule Always_LeadsTo_weaken, assumption+) |
|
346 |
apply (auto simp add: nat_eq_iff nat_less_iff) |
|
347 |
done |
|
348 |
||
349 |
lemma LessThan_bounded_induct: |
|
350 |
"!!l::nat. [| ALL m:(greaterThan l). F : (A Int f-`{m}) LeadsTo
|
|
351 |
((A Int f-`(lessThan m)) Un B) |] |
|
352 |
==> F : A LeadsTo ((A Int (f-`(atMost l))) Un B)" |
|
353 |
apply (simp only: Diff_eq [symmetric] vimage_Compl Compl_greaterThan [symmetric]) |
|
354 |
apply (rule wf_less_than [THEN Bounded_induct]) |
|
355 |
apply (simp (no_asm_simp)) |
|
356 |
done |
|
357 |
||
358 |
lemma GreaterThan_bounded_induct: |
|
359 |
"!!l::nat. [| ALL m:(lessThan l). F : (A Int f-`{m}) LeadsTo
|
|
360 |
((A Int f-`(greaterThan m)) Un B) |] |
|
361 |
==> F : A LeadsTo ((A Int (f-`(atLeast l))) Un B)" |
|
362 |
apply (rule_tac f = f and f1 = "%k. l - k" |
|
363 |
in wf_less_than [THEN wf_inv_image, THEN LeadsTo_wf_induct]) |
|
364 |
apply (simp add: inv_image_def Image_singleton, clarify) |
|
365 |
apply (case_tac "m<l") |
|
366 |
prefer 2 apply (blast intro: not_leE subset_imp_LeadsTo) |
|
367 |
apply (blast intro: LeadsTo_weaken_R diff_less_mono2) |
|
368 |
done |
|
369 |
||
370 |
||
371 |
(*** Completion: Binary and General Finite versions ***) |
|
372 |
||
373 |
lemma Completion: |
|
374 |
"[| F : A LeadsTo (A' Un C); F : A' Co (A' Un C); |
|
375 |
F : B LeadsTo (B' Un C); F : B' Co (B' Un C) |] |
|
376 |
==> F : (A Int B) LeadsTo ((A' Int B') Un C)" |
|
377 |
apply (simp (no_asm_use) add: LeadsTo_eq_leadsTo Constrains_eq_constrains Int_Un_distrib) |
|
378 |
apply (blast intro: completion leadsTo_weaken) |
|
379 |
done |
|
380 |
||
381 |
lemma Finite_completion_lemma: |
|
382 |
"finite I |
|
383 |
==> (ALL i:I. F : (A i) LeadsTo (A' i Un C)) --> |
|
384 |
(ALL i:I. F : (A' i) Co (A' i Un C)) --> |
|
385 |
F : (INT i:I. A i) LeadsTo ((INT i:I. A' i) Un C)" |
|
386 |
apply (erule finite_induct, auto) |
|
387 |
apply (rule Completion) |
|
388 |
prefer 4 |
|
389 |
apply (simp only: INT_simps [symmetric]) |
|
390 |
apply (rule Constrains_INT, auto) |
|
391 |
done |
|
392 |
||
393 |
lemma Finite_completion: |
|
394 |
"[| finite I; |
|
395 |
!!i. i:I ==> F : (A i) LeadsTo (A' i Un C); |
|
396 |
!!i. i:I ==> F : (A' i) Co (A' i Un C) |] |
|
397 |
==> F : (INT i:I. A i) LeadsTo ((INT i:I. A' i) Un C)" |
|
398 |
by (blast intro: Finite_completion_lemma [THEN mp, THEN mp]) |
|
399 |
||
400 |
lemma Stable_completion: |
|
401 |
"[| F : A LeadsTo A'; F : Stable A'; |
|
402 |
F : B LeadsTo B'; F : Stable B' |] |
|
403 |
==> F : (A Int B) LeadsTo (A' Int B')" |
|
404 |
apply (unfold Stable_def) |
|
405 |
apply (rule_tac C1 = "{}" in Completion [THEN LeadsTo_weaken_R])
|
|
406 |
apply (force+) |
|
407 |
done |
|
408 |
||
409 |
lemma Finite_stable_completion: |
|
410 |
"[| finite I; |
|
411 |
!!i. i:I ==> F : (A i) LeadsTo (A' i); |
|
412 |
!!i. i:I ==> F : Stable (A' i) |] |
|
413 |
==> F : (INT i:I. A i) LeadsTo (INT i:I. A' i)" |
|
414 |
apply (unfold Stable_def) |
|
415 |
apply (rule_tac C1 = "{}" in Finite_completion [THEN LeadsTo_weaken_R])
|
|
416 |
apply (simp_all (no_asm_simp)) |
|
417 |
apply blast+ |
|
418 |
done |
|
419 |
||
| 9685 | 420 |
|
| 4776 | 421 |
end |