author | paulson |
Tue, 21 Dec 1999 15:03:02 +0100 | |
changeset 8069 | 19b9f92ca503 |
parent 8067 | 225e3b45b766 |
child 8072 | 5b95377d7538 |
permissions | -rw-r--r-- |
8044 | 1 |
(* Title: HOL/UNITY/ELT |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
6 |
leadsTo strengthened with a specification of the allowable sets transient parts |
|
7 |
*) |
|
8 |
||
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
9 |
Goalw [givenBy_def] "givenBy id = UNIV"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
10 |
by Auto_tac; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
11 |
qed "givenBy_id"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
12 |
Addsimps [givenBy_id]; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
13 |
|
8044 | 14 |
Goalw [givenBy_def] "(givenBy v) = {A. ALL x:A. ALL y. v x = v y --> y: A}"; |
15 |
by Safe_tac; |
|
16 |
by (res_inst_tac [("x", "v `` ?u")] image_eqI 2); |
|
17 |
by Auto_tac; |
|
18 |
qed "givenBy_eq_all"; |
|
19 |
||
20 |
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}"; |
|
21 |
by (simp_tac (simpset() addsimps [givenBy_eq_all]) 1); |
|
22 |
by Safe_tac; |
|
23 |
by (res_inst_tac [("x", "%n. EX s. v s = n & s : ?A")] exI 1); |
|
24 |
by (Blast_tac 1); |
|
25 |
by Auto_tac; |
|
26 |
qed "givenBy_eq_Collect"; |
|
27 |
||
28 |
val prems = |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
29 |
Goal "(!!x y. [| x:A; v x = v y |] ==> y: A) ==> A: givenBy v"; |
8044 | 30 |
by (stac givenBy_eq_all 1); |
31 |
by (blast_tac (claset() addIs prems) 1); |
|
32 |
qed "givenByI"; |
|
33 |
||
34 |
Goalw [givenBy_def] "[| A: givenBy v; x:A; v x = v y |] ==> y: A"; |
|
35 |
by Auto_tac; |
|
36 |
qed "givenByD"; |
|
37 |
||
38 |
Goal "{} : givenBy v"; |
|
39 |
by (blast_tac (claset() addSIs [givenByI]) 1); |
|
40 |
qed "empty_mem_givenBy"; |
|
41 |
||
42 |
AddIffs [empty_mem_givenBy]; |
|
43 |
||
44 |
Goal "A: givenBy v ==> EX P. A = {s. P(v s)}"; |
|
45 |
by (res_inst_tac [("x", "%n. EX s. v s = n & s : A")] exI 1); |
|
46 |
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1); |
|
47 |
by (Blast_tac 1); |
|
48 |
qed "givenBy_imp_eq_Collect"; |
|
49 |
||
50 |
Goalw [givenBy_def] "EX P. A = {s. P(v s)} ==> A: givenBy v"; |
|
51 |
by (Best_tac 1); |
|
52 |
qed "eq_Collect_imp_givenBy"; |
|
53 |
||
54 |
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}"; |
|
55 |
by (blast_tac (claset() addIs [eq_Collect_imp_givenBy, |
|
56 |
givenBy_imp_eq_Collect]) 1); |
|
57 |
qed "givenBy_eq_eq_Collect"; |
|
58 |
||
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
59 |
(*preserving v preserves properties given by v*) |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
60 |
Goal "[| F : preserves v; D : givenBy v |] ==> F : stable D"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
61 |
by (force_tac (claset(), |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
62 |
simpset() addsimps [impOfSubs preserves_subset_stable, |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
63 |
givenBy_eq_Collect]) 1); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
64 |
qed "preserves_givenBy_imp_stable"; |
8044 | 65 |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
66 |
Goal "givenBy (w o v) <= givenBy v"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
67 |
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
68 |
by (Deepen_tac 0 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
69 |
qed "givenBy_o_subset"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
70 |
|
8044 | 71 |
|
72 |
(** Standard leadsTo rules **) |
|
73 |
||
74 |
Goalw [leadsETo_def] "[| F: A ensures B; A-B: CC |] ==> F : A leadsTo[CC] B"; |
|
75 |
by (blast_tac (claset() addIs [elt.Basis]) 1); |
|
76 |
qed "leadsETo_Basis"; |
|
77 |
||
78 |
Goalw [leadsETo_def] |
|
79 |
"[| F : A leadsTo[CC] B; F : B leadsTo[CC] C |] ==> F : A leadsTo[CC] C"; |
|
80 |
by (blast_tac (claset() addIs [elt.Trans]) 1); |
|
81 |
qed "leadsETo_Trans"; |
|
82 |
||
83 |
(*Useful with cancellation, disjunction*) |
|
84 |
Goal "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'"; |
|
85 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
|
86 |
qed "leadsETo_Un_duplicate"; |
|
87 |
||
88 |
Goal "F : A leadsTo[CC] (A' Un C Un C) ==> F : A leadsTo[CC] (A' Un C)"; |
|
89 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
|
90 |
qed "leadsETo_Un_duplicate2"; |
|
91 |
||
92 |
(*The Union introduction rule as we should have liked to state it*) |
|
93 |
val prems = Goalw [leadsETo_def] |
|
94 |
"(!!A. A : S ==> F : A leadsTo[CC] B) ==> F : (Union S) leadsTo[CC] B"; |
|
95 |
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1); |
|
96 |
qed "leadsETo_Union"; |
|
97 |
||
98 |
val prems = Goal |
|
99 |
"(!!i. i : I ==> F : (A i) leadsTo[CC] B) \ |
|
100 |
\ ==> F : (UN i:I. A i) leadsTo[CC] B"; |
|
101 |
by (stac (Union_image_eq RS sym) 1); |
|
102 |
by (blast_tac (claset() addIs leadsETo_Union::prems) 1); |
|
103 |
qed "leadsETo_UN"; |
|
104 |
||
105 |
(*The INDUCTION rule as we should have liked to state it*) |
|
106 |
val major::prems = Goalw [leadsETo_def] |
|
107 |
"[| F : za leadsTo[CC] zb; \ |
|
108 |
\ !!A B. [| F : A ensures B; A-B : CC |] ==> P A B; \ |
|
109 |
\ !!A B C. [| F : A leadsTo[CC] B; P A B; F : B leadsTo[CC] C; P B C |] \ |
|
110 |
\ ==> P A C; \ |
|
111 |
\ !!B S. ALL A:S. F : A leadsTo[CC] B & P A B ==> P (Union S) B \ |
|
112 |
\ |] ==> P za zb"; |
|
113 |
by (rtac (major RS CollectD RS elt.induct) 1); |
|
114 |
by (REPEAT (blast_tac (claset() addIs prems) 1)); |
|
115 |
qed "leadsETo_induct"; |
|
116 |
||
117 |
||
118 |
(** New facts involving leadsETo **) |
|
119 |
||
120 |
Goal "CC' <= CC ==> (A leadsTo[CC'] B) <= (A leadsTo[CC] B)"; |
|
121 |
by Safe_tac; |
|
122 |
by (etac leadsETo_induct 1); |
|
123 |
by (blast_tac (claset() addIs [leadsETo_Union]) 3); |
|
124 |
by (blast_tac (claset() addIs [leadsETo_Trans]) 2); |
|
125 |
by (blast_tac (claset() addIs [leadsETo_Basis]) 1); |
|
126 |
qed "leadsETo_mono"; |
|
127 |
||
128 |
val prems = Goalw [leadsETo_def] |
|
129 |
"(!!A. A : S ==> F : (A Int C) leadsTo[CC] B) ==> F : (Union S Int C) leadsTo[CC] B"; |
|
130 |
by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1); |
|
131 |
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1); |
|
132 |
qed "leadsETo_Union_Int"; |
|
133 |
||
134 |
(*Binary union introduction rule*) |
|
135 |
Goal "[| F : A leadsTo[CC] C; F : B leadsTo[CC] C |] ==> F : (A Un B) leadsTo[CC] C"; |
|
136 |
by (stac Un_eq_Union 1); |
|
137 |
by (blast_tac (claset() addIs [leadsETo_Union]) 1); |
|
138 |
qed "leadsETo_Un"; |
|
139 |
||
140 |
val prems = |
|
141 |
Goal "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B"; |
|
142 |
by (stac (UN_singleton RS sym) 1 THEN rtac leadsETo_UN 1); |
|
143 |
by (blast_tac (claset() addIs prems) 1); |
|
144 |
qed "single_leadsETo_I"; |
|
145 |
||
146 |
||
147 |
Goal "[| A<=B; {}:CC |] ==> F : A leadsTo[CC] B"; |
|
148 |
by (asm_simp_tac (simpset() addsimps [subset_imp_ensures RS leadsETo_Basis, |
|
149 |
Diff_eq_empty_iff RS iffD2]) 1); |
|
150 |
qed "subset_imp_leadsETo"; |
|
151 |
||
152 |
bind_thm ("empty_leadsETo", empty_subsetI RS subset_imp_leadsETo); |
|
153 |
Addsimps [empty_leadsETo]; |
|
154 |
||
155 |
||
156 |
(** Weakening laws all require {}:CC **) |
|
157 |
||
158 |
Goal "[| F : A leadsTo[CC] A'; A'<=B'; {}:CC |] ==> F : A leadsTo[CC] B'"; |
|
159 |
by (blast_tac (claset() addIs [subset_imp_leadsETo, leadsETo_Trans]) 1); |
|
160 |
qed "leadsETo_weaken_R"; |
|
161 |
||
162 |
Goal "[| F : A leadsTo[CC] A'; B<=A; {}:CC |] ==> F : B leadsTo[CC] A'"; |
|
163 |
by (blast_tac (claset() addIs [leadsETo_Trans, subset_imp_leadsETo]) 1); |
|
164 |
qed_spec_mp "leadsETo_weaken_L"; |
|
165 |
||
166 |
(*Distributes over binary unions*) |
|
167 |
Goal "{} : CC ==> \ |
|
168 |
\ F : (A Un B) leadsTo[CC] C = (F : A leadsTo[CC] C & F : B leadsTo[CC] C)"; |
|
169 |
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken_L]) 1); |
|
170 |
qed "leadsETo_Un_distrib"; |
|
171 |
||
172 |
Goal "{} : CC ==> \ |
|
173 |
\ F : (UN i:I. A i) leadsTo[CC] B = (ALL i : I. F : (A i) leadsTo[CC] B)"; |
|
174 |
by (blast_tac (claset() addIs [leadsETo_UN, leadsETo_weaken_L]) 1); |
|
175 |
qed "leadsETo_UN_distrib"; |
|
176 |
||
177 |
Goal "{} : CC \ |
|
178 |
\ ==> F : (Union S) leadsTo[CC] B = (ALL A : S. F : A leadsTo[CC] B)"; |
|
179 |
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_L]) 1); |
|
180 |
qed "leadsETo_Union_distrib"; |
|
181 |
||
182 |
Goal "[| F : A leadsTo[CC'] A'; B<=A; A'<=B'; CC' <= CC; {}:CC |] \ |
|
183 |
\ ==> F : B leadsTo[CC] B'"; |
|
184 |
by (dtac (impOfSubs leadsETo_mono) 1); |
|
185 |
by (assume_tac 1); |
|
186 |
by (blast_tac (claset() addIs [leadsETo_weaken_R, leadsETo_weaken_L, |
|
187 |
leadsETo_Trans]) 1); |
|
188 |
qed "leadsETo_weaken"; |
|
189 |
||
190 |
Goal "[| F : A leadsTo[CC] A'; CC <= givenBy v |] \ |
|
191 |
\ ==> F : A leadsTo[givenBy v] A'"; |
|
192 |
by (blast_tac (claset() addIs [empty_mem_givenBy, leadsETo_weaken]) 1); |
|
193 |
qed "leadsETo_givenBy"; |
|
194 |
||
195 |
||
196 |
(*Set difference*) |
|
197 |
Goal "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C; {}:CC |] \ |
|
198 |
\ ==> F : A leadsTo[CC] C"; |
|
199 |
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken]) 1); |
|
200 |
qed "leadsETo_Diff"; |
|
201 |
||
202 |
||
203 |
(** Meta or object quantifier ??? |
|
204 |
see ball_constrains_UN in UNITY.ML***) |
|
205 |
||
206 |
val prems = goal thy |
|
207 |
"[| !! i. i:I ==> F : (A i) leadsTo[CC] (A' i); {}:CC |] \ |
|
208 |
\ ==> F : (UN i:I. A i) leadsTo[CC] (UN i:I. A' i)"; |
|
209 |
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1); |
|
210 |
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_R] |
|
211 |
addIs prems) 1); |
|
212 |
qed "leadsETo_UN_UN"; |
|
213 |
||
214 |
(*Binary union version*) |
|
215 |
Goal "[| F : A leadsTo[CC] A'; F : B leadsTo[CC] B'; {}:CC |] \ |
|
216 |
\ ==> F : (A Un B) leadsTo[CC] (A' Un B')"; |
|
217 |
by (blast_tac (claset() addIs [leadsETo_Un, |
|
218 |
leadsETo_weaken_R]) 1); |
|
219 |
qed "leadsETo_Un_Un"; |
|
220 |
||
221 |
||
222 |
(** The cancellation law **) |
|
223 |
||
224 |
Goal "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B'; {}:CC |] \ |
|
225 |
\ ==> F : A leadsTo[CC] (A' Un B')"; |
|
226 |
by (blast_tac (claset() addIs [leadsETo_Un_Un, |
|
227 |
subset_imp_leadsETo, leadsETo_Trans]) 1); |
|
228 |
qed "leadsETo_cancel2"; |
|
229 |
||
230 |
Goal "[| F : A leadsTo[CC] (A' Un B); F : (B-A') leadsTo[CC] B'; {}:CC |] \ |
|
231 |
\ ==> F : A leadsTo[CC] (A' Un B')"; |
|
232 |
by (rtac leadsETo_cancel2 1); |
|
233 |
by (assume_tac 2); |
|
234 |
by (ALLGOALS Asm_simp_tac); |
|
235 |
qed "leadsETo_cancel_Diff2"; |
|
236 |
||
237 |
Goal "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B'; {}:CC |] \ |
|
238 |
\ ==> F : A leadsTo[CC] (B' Un A')"; |
|
239 |
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1); |
|
240 |
by (blast_tac (claset() addSIs [leadsETo_cancel2]) 1); |
|
241 |
qed "leadsETo_cancel1"; |
|
242 |
||
243 |
Goal "[| F : A leadsTo[CC] (B Un A'); F : (B-A') leadsTo[CC] B'; {}:CC |] \ |
|
244 |
\ ==> F : A leadsTo[CC] (B' Un A')"; |
|
245 |
by (rtac leadsETo_cancel1 1); |
|
246 |
by (assume_tac 2); |
|
247 |
by (ALLGOALS Asm_simp_tac); |
|
248 |
qed "leadsETo_cancel_Diff1"; |
|
249 |
||
250 |
||
251 |
(** The impossibility law **) |
|
252 |
||
253 |
Goal "F : A leadsTo[CC] B ==> B={} --> A={}"; |
|
254 |
by (etac leadsETo_induct 1); |
|
255 |
by (ALLGOALS Asm_simp_tac); |
|
256 |
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]); |
|
257 |
by (Blast_tac 1); |
|
258 |
val lemma = result() RS mp; |
|
259 |
||
260 |
Goal "F : A leadsTo[CC] {} ==> A={}"; |
|
261 |
by (blast_tac (claset() addSIs [lemma]) 1); |
|
262 |
qed "leadsETo_empty"; |
|
263 |
||
264 |
||
265 |
(** PSP: Progress-Safety-Progress **) |
|
266 |
||
267 |
(*Special case of PSP: Misra's "stable conjunction"*) |
|
268 |
Goalw [stable_def] |
|
269 |
"[| F : A leadsTo[CC] A'; F : stable B; ALL C:CC. C Int B : CC |] \ |
|
270 |
\ ==> F : (A Int B) leadsTo[CC] (A' Int B)"; |
|
271 |
by (etac leadsETo_induct 1); |
|
272 |
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3); |
|
273 |
by (blast_tac (claset() addIs [leadsETo_Trans]) 2); |
|
274 |
by (rtac leadsETo_Basis 1); |
|
275 |
by (asm_full_simp_tac |
|
276 |
(simpset() addsimps [ensures_def, |
|
277 |
Diff_Int_distrib2 RS sym, Int_Un_distrib2 RS sym]) 1); |
|
278 |
by (asm_simp_tac (simpset() addsimps [Diff_Int_distrib2 RS sym]) 2); |
|
279 |
by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1); |
|
280 |
qed "e_psp_stable"; |
|
281 |
||
282 |
Goal "[| F : A leadsTo[CC] A'; F : stable B; ALL C:CC. C Int B : CC |] \ |
|
283 |
\ ==> F : (B Int A) leadsTo[CC] (B Int A')"; |
|
284 |
by (asm_simp_tac (simpset() addsimps e_psp_stable::Int_ac) 1); |
|
285 |
qed "e_psp_stable2"; |
|
286 |
||
287 |
Goal "[| F : A leadsTo[CC] A'; F : B co B'; \ |
|
288 |
\ ALL C:CC. C Int B Int B' : CC; {}:CC |] \ |
|
289 |
\ ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))"; |
|
290 |
by (etac leadsETo_induct 1); |
|
291 |
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3); |
|
292 |
(*Transitivity case has a delicate argument involving "cancellation"*) |
|
293 |
by (rtac leadsETo_Un_duplicate2 2); |
|
294 |
by (etac leadsETo_cancel_Diff1 2); |
|
295 |
by (assume_tac 3); |
|
296 |
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2); |
|
297 |
by (blast_tac (claset() addIs [leadsETo_weaken_L] |
|
298 |
addDs [constrains_imp_subset]) 2); |
|
299 |
(*Basis case*) |
|
300 |
by (rtac leadsETo_Basis 1); |
|
301 |
by (blast_tac (claset() addIs [psp_ensures]) 1); |
|
302 |
by (subgoal_tac "A Int B' - (Ba Int B Un (B' - B)) = (A - Ba) Int B Int B'" 1); |
|
303 |
by Auto_tac; |
|
304 |
qed "e_psp"; |
|
305 |
||
306 |
Goal "[| F : A leadsTo[CC] A'; F : B co B'; \ |
|
307 |
\ ALL C:CC. C Int B Int B' : CC; {}:CC |] \ |
|
308 |
\ ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))"; |
|
309 |
by (asm_full_simp_tac (simpset() addsimps e_psp::Int_ac) 1); |
|
310 |
qed "e_psp2"; |
|
311 |
||
312 |
||
313 |
(*** Special properties involving the parameter [CC] ***) |
|
314 |
||
315 |
(*??IS THIS NEEDED?? or is it just an example of what's provable??*) |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
316 |
Goal "[| F: (A leadsTo[givenBy v] B); G : preserves v; \ |
8044 | 317 |
\ F Join G : stable C |] \ |
318 |
\ ==> F Join G : ((C Int A) leadsTo[(%D. C Int D) `` givenBy v] B)"; |
|
319 |
by (etac leadsETo_induct 1); |
|
320 |
by (stac Int_Union 3); |
|
321 |
by (blast_tac (claset() addIs [leadsETo_UN]) 3); |
|
322 |
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, |
|
323 |
leadsETo_Trans]) 2); |
|
324 |
by (rtac leadsETo_Basis 1); |
|
325 |
by (auto_tac (claset(), |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
326 |
simpset() addsimps [Int_Diff, ensures_def, |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
327 |
givenBy_eq_Collect, Join_stable, |
8044 | 328 |
Join_constrains, Join_transient])); |
329 |
by (blast_tac (claset() addIs [transient_strengthen]) 3); |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
330 |
by (ALLGOALS (dres_inst_tac [("P1","P")] (impOfSubs preserves_subset_stable))); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
331 |
by (rewtac stable_def); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
332 |
by (blast_tac (claset() addSEs [equalityE] |
8044 | 333 |
addIs [constrains_Int RS constrains_weaken]) 2); |
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
334 |
by (blast_tac (claset() addSEs [equalityE] |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
335 |
addIs [constrains_Int RS constrains_weaken]) 1); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
336 |
qed "gen_leadsETo_imp_Join_leadsETo"; |
8044 | 337 |
|
338 |
(*useful??*) |
|
339 |
Goal "[| F Join G : (A leadsTo[CC] B); ALL C:CC. G : stable C |] \ |
|
340 |
\ ==> F: (A leadsTo[CC] B)"; |
|
341 |
by (etac leadsETo_induct 1); |
|
342 |
by (blast_tac (claset() addIs [leadsETo_Union]) 3); |
|
343 |
by (blast_tac (claset() addIs [leadsETo_Trans]) 2); |
|
344 |
by (rtac leadsETo_Basis 1); |
|
345 |
by (case_tac "A <= B" 1); |
|
346 |
by (etac subset_imp_ensures 1); |
|
347 |
by (auto_tac (claset() addIs [constrains_weaken], |
|
348 |
simpset() addsimps [stable_def, ensures_def, |
|
349 |
Join_constrains, Join_transient])); |
|
350 |
by (REPEAT (thin_tac "?F : ?A co ?B" 1)); |
|
351 |
by (etac transientE 1); |
|
352 |
by (rewtac constrains_def); |
|
353 |
by (blast_tac (claset() addSDs [bspec]) 1); |
|
354 |
qed "Join_leadsETo_stable_imp_leadsETo"; |
|
355 |
||
356 |
||
8067 | 357 |
(**** Relationship with traditional "leadsTo", strong & weak ****) |
358 |
||
359 |
(** strong **) |
|
360 |
||
361 |
Goal "(A leadsTo[CC] B) <= (A leadsTo B)"; |
|
362 |
by Safe_tac; |
|
363 |
by (etac leadsETo_induct 1); |
|
364 |
by (blast_tac (claset() addIs [leadsTo_Union]) 3); |
|
365 |
by (blast_tac (claset() addIs [leadsTo_Trans]) 2); |
|
366 |
by (blast_tac (claset() addIs [leadsTo_Basis]) 1); |
|
367 |
qed "leadsETo_subset_leadsTo"; |
|
368 |
||
369 |
Goal "(A leadsTo[UNIV] B) = (A leadsTo B)"; |
|
370 |
by Safe_tac; |
|
371 |
by (etac (impOfSubs leadsETo_subset_leadsTo) 1); |
|
372 |
(*right-to-left case*) |
|
373 |
by (etac leadsTo_induct 1); |
|
374 |
by (blast_tac (claset() addIs [leadsETo_Union]) 3); |
|
375 |
by (blast_tac (claset() addIs [leadsETo_Trans]) 2); |
|
376 |
by (blast_tac (claset() addIs [leadsETo_Basis]) 1); |
|
377 |
qed "leadsETo_UNIV_eq_leadsTo"; |
|
378 |
||
379 |
(** weak **) |
|
380 |
||
381 |
Goalw [LeadsETo_def, LeadsTo_def] "(A LeadsTo[CC] B) <= (A LeadsTo B)"; |
|
382 |
by (blast_tac (claset() addIs [impOfSubs leadsETo_subset_leadsTo]) 1); |
|
383 |
qed "LeadsETo_subset_LeadsTo"; |
|
384 |
||
385 |
(*Postcondition can be strengthened to (reachable F Int B) *) |
|
386 |
Goal "F : A ensures B ==> F : (reachable F Int A) ensures B"; |
|
387 |
by (rtac (stable_ensures_Int RS ensures_weaken_R) 1); |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
388 |
by Auto_tac; |
8067 | 389 |
qed "reachable_ensures"; |
390 |
||
391 |
Goal "F : A leadsTo B ==> F : (reachable F Int A) leadsTo[Pow(reachable F)] B"; |
|
392 |
by (etac leadsTo_induct 1); |
|
393 |
by (stac Int_Union 3); |
|
394 |
by (blast_tac (claset() addIs [leadsETo_UN]) 3); |
|
395 |
by (blast_tac (claset() addDs [e_psp_stable2] |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
396 |
addIs [leadsETo_Trans, leadsETo_weaken_L]) 2); |
8067 | 397 |
by (blast_tac (claset() addIs [reachable_ensures, leadsETo_Basis]) 1); |
398 |
val lemma = result(); |
|
399 |
||
400 |
Goal "(A LeadsTo[UNIV] B) = (A LeadsTo B)"; |
|
401 |
by Safe_tac; |
|
402 |
by (etac (impOfSubs LeadsETo_subset_LeadsTo) 1); |
|
403 |
(*right-to-left case*) |
|
404 |
by (rewrite_goals_tac [LeadsETo_def, LeadsTo_def]); |
|
405 |
by (fast_tac (claset() addEs [lemma RS leadsETo_weaken]) 1); |
|
406 |
qed "LeadsETo_UNIV_eq_LeadsTo"; |
|
407 |
||
8044 | 408 |
|
409 |
(**** EXTEND/PROJECT PROPERTIES ****) |
|
410 |
||
411 |
Open_locale "Extend"; |
|
412 |
||
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
413 |
Goal "givenBy (v o f) = extend_set h `` (givenBy v)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
414 |
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
415 |
by (Deepen_tac 0 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
416 |
qed "givenBy_o_eq_extend_set"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
417 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
418 |
Goal "givenBy f = range (extend_set h)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
419 |
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
420 |
by (Deepen_tac 0 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
421 |
qed "givenBy_eq_extend_set"; |
8044 | 422 |
|
423 |
Goal "D : givenBy v ==> extend_set h D : givenBy (v o f)"; |
|
424 |
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1); |
|
425 |
by (Blast_tac 1); |
|
426 |
qed "extend_set_givenBy_I"; |
|
427 |
||
428 |
||
429 |
Goal "F : A leadsTo[CC] B \ |
|
430 |
\ ==> extend h F : (extend_set h A) leadsTo[extend_set h `` CC] \ |
|
431 |
\ (extend_set h B)"; |
|
432 |
by (etac leadsETo_induct 1); |
|
433 |
by (asm_simp_tac (simpset() addsimps [leadsETo_UN, extend_set_Union]) 3); |
|
434 |
by (blast_tac (claset() addIs [leadsETo_Trans]) 2); |
|
435 |
by (asm_simp_tac (simpset() addsimps [leadsETo_Basis, extend_ensures, |
|
436 |
extend_set_Diff_distrib RS sym]) 1); |
|
437 |
qed "leadsETo_imp_extend_leadsETo"; |
|
438 |
||
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
439 |
(*MOVE to Extend.ML?*) |
8044 | 440 |
Goal "A Int extend_set h ((project_set h A) Int B) = A Int extend_set h B"; |
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
441 |
by (auto_tac (claset(), simpset() addsimps [split_extended_all])); |
8044 | 442 |
qed "Int_extend_set_lemma"; |
443 |
||
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
444 |
(*MOVE to extend.ML??*) |
8044 | 445 |
Goal "G : C co B ==> project h C G : project_set h C co project_set h B"; |
446 |
by (full_simp_tac (simpset() addsimps [constrains_def, project_def, |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
447 |
project_act_def]) 1); |
8044 | 448 |
by (Blast_tac 1); |
449 |
qed "project_constrains_project_set"; |
|
450 |
||
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
451 |
(*MOVE to extend.ML??*) |
8044 | 452 |
Goal "G : stable C ==> project h C G : stable (project_set h C)"; |
453 |
by (asm_full_simp_tac (simpset() addsimps [stable_def, |
|
454 |
project_constrains_project_set]) 1); |
|
455 |
qed "project_stable_project_set"; |
|
456 |
||
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
457 |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
458 |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
459 |
(*NOT USED, but analogous to preserves_project_transient_empty in Project.ML*) |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
460 |
Goal "[| G : preserves (v o f); project h C G : transient D; \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
461 |
\ D : givenBy v |] ==> D={}"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
462 |
by (rtac stable_transient_empty 1); |
8044 | 463 |
by (assume_tac 2); |
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
464 |
(*If addIs then PROOF FAILED at depth 2*) |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
465 |
by (blast_tac (claset() addSIs [preserves_givenBy_imp_stable, |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
466 |
project_preserves_I]) 1); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
467 |
result(); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
468 |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
469 |
(*Generalizes the version proved in Project.ML*) |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
470 |
Goal "[| G : preserves (v o f); \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
471 |
\ project h C G : transient (C' Int D); \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
472 |
\ project h C G : stable C'; D : givenBy v |] \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
473 |
\ ==> C' Int D = {}"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
474 |
by (rtac stable_transient_empty 1); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
475 |
by (assume_tac 2); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
476 |
(*If addIs then PROOF FAILED at depth 3*) |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
477 |
by (blast_tac (claset() addSIs [stable_Int, preserves_givenBy_imp_stable, |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
478 |
project_preserves_I]) 1); |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
479 |
qed "preserves_o_project_transient_empty"; |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
480 |
|
8044 | 481 |
|
482 |
(*This version's stronger in the "ensures" precondition |
|
483 |
BUT there's no ensures_weaken_L*) |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
484 |
Goal "[| project h C G ~: transient (project_set h C Int (A-B)) | \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
485 |
\ project_set h C Int (A - B) = {}; \ |
8044 | 486 |
\ extend h F Join G : stable C; \ |
487 |
\ F Join project h C G : (project_set h C Int A) ensures B |] \ |
|
488 |
\ ==> extend h F Join G : (C Int extend_set h A) ensures (extend_set h B)"; |
|
489 |
by (stac (Int_extend_set_lemma RS sym) 1); |
|
490 |
by (rtac Join_project_ensures 1); |
|
491 |
by (auto_tac (claset(), simpset() addsimps [Int_Diff])); |
|
492 |
qed "Join_project_ensures_strong"; |
|
493 |
||
494 |
Goal "[| extend h F Join G : stable C; \ |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
495 |
\ F Join project h C G : (project_set h C Int A) leadsTo[(%D. project_set h C Int D)``givenBy v] B; \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
496 |
\ G : preserves (v o f) |] \ |
8044 | 497 |
\ ==> extend h F Join G : \ |
498 |
\ (C Int extend_set h (project_set h C Int A)) \ |
|
499 |
\ leadsTo[(%D. C Int extend_set h D)``givenBy v] (extend_set h B)"; |
|
500 |
by (etac leadsETo_induct 1); |
|
501 |
by (asm_simp_tac (simpset() delsimps UN_simps |
|
502 |
addsimps [Int_UN_distrib, leadsETo_UN, extend_set_Union]) 3); |
|
503 |
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, |
|
504 |
leadsETo_Trans]) 2); |
|
505 |
by (Clarify_tac 1); |
|
506 |
by (rtac leadsETo_Basis 1); |
|
507 |
by (etac rev_image_eqI 2); |
|
508 |
by (asm_simp_tac (simpset() addsimps [Int_Diff, Int_extend_set_lemma, |
|
509 |
extend_set_Diff_distrib RS sym]) 2); |
|
510 |
by (rtac Join_project_ensures_strong 1); |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
511 |
by (auto_tac (claset() addDs [preserves_o_project_transient_empty] |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
512 |
addIs [project_stable_project_set], |
8044 | 513 |
simpset() addsimps [Int_left_absorb, Join_stable])); |
514 |
by (asm_simp_tac |
|
515 |
(simpset() addsimps [stable_ensures_Int RS ensures_weaken_R, |
|
516 |
Int_lower2, project_stable_project_set, |
|
517 |
Join_stable, extend_stable_project_set]) 1); |
|
518 |
val lemma = result(); |
|
519 |
||
520 |
Goal "[| extend h F Join G : stable C; \ |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
521 |
\ F Join project h C G : \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
522 |
\ (project_set h C Int A) \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
523 |
\ leadsTo[(%D. project_set h C Int D)``givenBy v] B; \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
524 |
\ G : preserves (v o f) |] \ |
8044 | 525 |
\ ==> extend h F Join G : (C Int extend_set h A) \ |
526 |
\ leadsTo[(%D. C Int extend_set h D)``givenBy v] (extend_set h B)"; |
|
527 |
by (rtac (lemma RS leadsETo_weaken) 1); |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
528 |
by (auto_tac (claset(), |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
529 |
simpset() addsimps [split_extended_all])); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
530 |
qed "project_leadsETo_D_lemma"; |
8044 | 531 |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
532 |
Goal "[| F Join project h UNIV G : A leadsTo[givenBy v] B; \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
533 |
\ G : preserves (v o f) |] \ |
8044 | 534 |
\ ==> extend h F Join G : (extend_set h A) \ |
535 |
\ leadsTo[givenBy (v o f)] (extend_set h B)"; |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
536 |
by (rtac (make_elim project_leadsETo_D_lemma) 1); |
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
537 |
by (stac stable_UNIV 1); |
8044 | 538 |
by Auto_tac; |
539 |
by (etac leadsETo_givenBy 1); |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
540 |
by (rtac (givenBy_o_eq_extend_set RS equalityD2) 1); |
8044 | 541 |
qed "project_leadsETo_D"; |
542 |
||
543 |
Goal "[| F Join project h (reachable (extend h F Join G)) G \ |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
544 |
\ : A LeadsTo[givenBy v] B; \ |
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
545 |
\ G : preserves (v o f) |] \ |
8044 | 546 |
\ ==> extend h F Join G : \ |
547 |
\ (extend_set h A) LeadsTo[givenBy (v o f)] (extend_set h B)"; |
|
548 |
by (rtac (make_elim (subset_refl RS stable_reachable RS |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
549 |
project_leadsETo_D_lemma)) 1); |
8044 | 550 |
by (auto_tac (claset(), |
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
551 |
simpset() addsimps [LeadsETo_def])); |
8044 | 552 |
by (asm_full_simp_tac |
553 |
(simpset() addsimps [project_set_reachable_extend_eq RS sym]) 1); |
|
554 |
by (etac (impOfSubs leadsETo_mono) 1); |
|
555 |
by (blast_tac (claset() addIs [extend_set_givenBy_I]) 1); |
|
556 |
qed "project_LeadsETo_D"; |
|
557 |
||
558 |
Goalw [extending_def] |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
559 |
"extending (v o f) (%G. UNIV) h F \ |
8044 | 560 |
\ (extend_set h A leadsTo[givenBy (v o f)] extend_set h B) \ |
561 |
\ (A leadsTo[givenBy v] B)"; |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
562 |
by (auto_tac (claset(), simpset() addsimps [project_leadsETo_D])); |
8044 | 563 |
qed "extending_leadsETo"; |
564 |
||
565 |
||
566 |
Goalw [extending_def] |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
567 |
"extending (v o f) (%G. reachable (extend h F Join G)) h F \ |
8044 | 568 |
\ (extend_set h A LeadsTo[givenBy (v o f)] extend_set h B) \ |
569 |
\ (A LeadsTo[givenBy v] B)"; |
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8044
diff
changeset
|
570 |
by (blast_tac (claset() addIs [project_LeadsETo_D]) 1); |
8044 | 571 |
qed "extending_LeadsETo"; |
572 |
||
573 |
||
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
574 |
(*** leadsETo in the precondition ***) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
575 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
576 |
Goalw [transient_def] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
577 |
"[| G : transient (C Int extend_set h A); G : stable C |] \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
578 |
\ ==> project h C G : transient (project_set h C Int A)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
579 |
by (auto_tac (claset(), simpset() addsimps [Domain_project_act])); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
580 |
by (subgoal_tac "act ^^ (C Int extend_set h A) <= - extend_set h A" 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
581 |
by (asm_full_simp_tac |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
582 |
(simpset() addsimps [stable_def, constrains_def]) 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
583 |
by (Blast_tac 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
584 |
(*back to main goal*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
585 |
by (thin_tac "?AA <= -C Un ?BB" 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
586 |
by (ball_tac 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
587 |
by (asm_full_simp_tac |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
588 |
(simpset() addsimps [extend_set_def, project_act_def]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
589 |
by (Blast_tac 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
590 |
qed "transient_extend_set_imp_project_transient"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
591 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
592 |
(*converse might hold too?*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
593 |
Goalw [transient_def] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
594 |
"project h C (extend h F) : transient (project_set h C Int D) \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
595 |
\ ==> F : transient (project_set h C Int D)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
596 |
by (auto_tac (claset(), simpset() addsimps [Domain_project_act])); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
597 |
by (rtac bexI 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
598 |
by (assume_tac 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
599 |
by Auto_tac; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
600 |
by (rewtac extend_act_def); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
601 |
by (Blast_tac 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
602 |
qed "project_extend_transient_D"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
603 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
604 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
605 |
Goal "[| extend h F : stable C; G : stable C; \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
606 |
\ extend h F Join G : A ensures B; A-B = C Int extend_set h D |] \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
607 |
\ ==> F Join project h C G \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
608 |
\ : (project_set h C Int project_set h A) ensures (project_set h B)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
609 |
by (asm_full_simp_tac |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
610 |
(simpset() addsimps [ensures_def, Join_constrains, project_constrains, |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
611 |
Join_transient, extend_transient]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
612 |
by (Clarify_tac 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
613 |
by (REPEAT_FIRST (rtac conjI)); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
614 |
(*first subgoal*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
615 |
by (blast_tac (claset() addIs [extend_stable_project_set RS stableD RS |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
616 |
constrains_Int RS constrains_weaken] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
617 |
addSDs [extend_constrains_project_set] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
618 |
addSEs [equalityE]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
619 |
(*2nd subgoal*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
620 |
by (etac (stableD RS constrains_Int RS constrains_weaken) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
621 |
by (assume_tac 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
622 |
by (Blast_tac 3); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
623 |
by (full_simp_tac (simpset() addsimps [extend_set_Int_distrib, |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
624 |
extend_set_Un_distrib]) 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
625 |
by (blast_tac (claset() addSIs [impOfSubs extend_set_project_set]) 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
626 |
by (full_simp_tac (simpset() addsimps [extend_set_def]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
627 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
628 |
(*The transient part*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
629 |
by Auto_tac; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
630 |
by (force_tac (claset() addSEs [equalityE] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
631 |
addIs [transient_extend_set_imp_project_transient RS |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
632 |
transient_strengthen], |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
633 |
simpset()) 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
634 |
by (full_simp_tac (simpset() addsimps [Int_Diff]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
635 |
by (force_tac (claset() addSEs [equalityE] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
636 |
addIs [transient_extend_set_imp_project_transient RS |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
637 |
project_extend_transient_D RS transient_strengthen], |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
638 |
simpset()) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
639 |
qed "ensures_extend_set_imp_project_ensures"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
640 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
641 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
642 |
(*Lemma for the Trans case*) |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
643 |
Goal "[| extend h F Join G : stable C; \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
644 |
\ F Join project h C G \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
645 |
\ : project_set h C Int project_set h A leadsTo project_set h B |] \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
646 |
\ ==> F Join project h C G \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
647 |
\ : project_set h C Int project_set h A leadsTo \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
648 |
\ project_set h C Int project_set h B"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
649 |
by (rtac (psp_stable2 RS leadsTo_weaken_L) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
650 |
by (auto_tac (claset(), |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
651 |
simpset() addsimps [project_stable_project_set, Join_stable, |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
652 |
extend_stable_project_set])); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
653 |
val lemma = result(); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
654 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
655 |
Goal "[| extend h F Join G : stable C; \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
656 |
\ extend h F Join G : \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
657 |
\ (C Int A) leadsTo[(%D. C Int D)``givenBy f] B |] \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
658 |
\ ==> F Join project h C G \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
659 |
\ : (project_set h C Int project_set h (C Int A)) leadsTo (project_set h B)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
660 |
by (etac leadsETo_induct 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
661 |
by (asm_simp_tac (HOL_ss addsimps [Int_UN_distrib, project_set_Union]) 3); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
662 |
by (blast_tac (claset() addIs [leadsTo_UN]) 3); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
663 |
by (blast_tac (claset() addIs [leadsTo_Trans, lemma]) 2); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
664 |
by (asm_full_simp_tac |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
665 |
(simpset() addsimps [givenBy_eq_extend_set, Join_stable]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
666 |
by (rtac leadsTo_Basis 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
667 |
by (blast_tac (claset() addIs [leadsTo_Basis, |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
668 |
ensures_extend_set_imp_project_ensures]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
669 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
670 |
qed "project_leadsETo_I_lemma"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
671 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
672 |
Goal "extend h F Join G : (extend_set h A) leadsTo[givenBy f] (extend_set h B)\ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
673 |
\ ==> F Join project h UNIV G : A leadsTo B"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
674 |
by (rtac (project_leadsETo_I_lemma RS leadsTo_weaken) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
675 |
by Auto_tac; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
676 |
qed "project_leadsETo_I"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
677 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
678 |
Goal "extend h F Join G : (extend_set h A) LeadsTo[givenBy f] (extend_set h B)\ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
679 |
\ ==> F Join project h (reachable (extend h F Join G)) G \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
680 |
\ : A LeadsTo B"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
681 |
by (full_simp_tac (simpset() addsimps [LeadsTo_def, LeadsETo_def]) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
682 |
by (rtac (project_leadsETo_I_lemma RS leadsTo_weaken) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
683 |
by (auto_tac (claset(), |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
684 |
simpset() addsimps [project_set_reachable_extend_eq RS sym])); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
685 |
qed "project_LeadsETo_I"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
686 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
687 |
Goalw [projecting_def] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
688 |
"projecting (%G. UNIV) h F \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
689 |
\ (extend_set h A leadsTo[givenBy f] extend_set h B) \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
690 |
\ (A leadsTo B)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
691 |
by (force_tac (claset() addDs [project_leadsETo_I], simpset()) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
692 |
qed "projecting_leadsTo"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
693 |
|
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
694 |
Goalw [projecting_def] |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
695 |
"projecting (%G. reachable (extend h F Join G)) h F \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
696 |
\ (extend_set h A LeadsTo[givenBy f] extend_set h B) \ |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
697 |
\ (A LeadsTo B)"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
698 |
by (force_tac (claset() addDs [project_LeadsETo_I], simpset()) 1); |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
699 |
qed "projecting_LeadsTo"; |
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8067
diff
changeset
|
700 |
|
8044 | 701 |
Close_locale "Extend"; |
702 |
||
703 |