author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17233 | 41eee2e7b465 |
child 19360 | f47412f922ab |
permissions | -rw-r--r-- |
4559 | 1 |
(* Title: HOLCF/IOA/meta_theory/Abstraction.thy |
2 |
ID: $Id$ |
|
12218 | 3 |
Author: Olaf Müller |
4559 | 4 |
*) |
5 |
||
6 |
section "cex_abs"; |
|
7 |
||
8 |
||
9 |
(* ---------------------------------------------------------------- *) |
|
10 |
(* cex_abs *) |
|
11 |
(* ---------------------------------------------------------------- *) |
|
12 |
||
5068 | 13 |
Goal "cex_abs f (s,UU) = (f s, UU)"; |
4559 | 14 |
by (simp_tac (simpset() addsimps [cex_abs_def]) 1); |
15 |
qed"cex_abs_UU"; |
|
16 |
||
5068 | 17 |
Goal "cex_abs f (s,nil) = (f s, nil)"; |
4559 | 18 |
by (simp_tac (simpset() addsimps [cex_abs_def]) 1); |
19 |
qed"cex_abs_nil"; |
|
20 |
||
5068 | 21 |
Goal "cex_abs f (s,(a,t)>>ex) = (f s, (a,f t) >> (snd (cex_abs f (t,ex))))"; |
4559 | 22 |
by (simp_tac (simpset() addsimps [cex_abs_def]) 1); |
23 |
qed"cex_abs_cons"; |
|
24 |
||
25 |
Addsimps [cex_abs_UU, cex_abs_nil, cex_abs_cons]; |
|
26 |
||
27 |
||
28 |
||
29 |
section "lemmas"; |
|
30 |
||
31 |
(* ---------------------------------------------------------------- *) |
|
32 |
(* Lemmas *) |
|
33 |
(* ---------------------------------------------------------------- *) |
|
34 |
||
5068 | 35 |
Goal "temp_weakening Q P h = (! ex. (ex |== P) --> (cex_abs h ex |== Q))"; |
4559 | 36 |
by (simp_tac (simpset() addsimps [temp_weakening_def,temp_strengthening_def, |
37 |
NOT_def,temp_sat_def,satisfies_def]) 1); |
|
5132 | 38 |
by Auto_tac; |
4559 | 39 |
qed"temp_weakening_def2"; |
40 |
||
5068 | 41 |
Goal "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))"; |
4559 | 42 |
by (simp_tac (simpset() addsimps [state_weakening_def,state_strengthening_def, |
43 |
NOT_def]) 1); |
|
5132 | 44 |
by Auto_tac; |
4559 | 45 |
qed"state_weakening_def2"; |
46 |
||
47 |
||
48 |
section "Abstraction Rules for Properties"; |
|
49 |
||
50 |
(* ---------------------------------------------------------------- *) |
|
51 |
(* Abstraction Rules for Properties *) |
|
52 |
(* ---------------------------------------------------------------- *) |
|
53 |
||
54 |
||
5068 | 55 |
Goalw [cex_abs_def] |
6161 | 56 |
"[| is_abstraction h C A |] ==>\ |
4559 | 57 |
\ !s. reachable C s & is_exec_frag C (s,xs) \ |
58 |
\ --> is_exec_frag A (cex_abs h (s,xs))"; |
|
59 |
||
60 |
by (Asm_full_simp_tac 1); |
|
61 |
by (pair_induct_tac "xs" [is_exec_frag_def] 1); |
|
62 |
(* main case *) |
|
63 |
by (safe_tac set_cs); |
|
64 |
by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def])1); |
|
65 |
by (forward_tac [reachable.reachable_n] 1); |
|
66 |
by (assume_tac 1); |
|
67 |
by (Asm_full_simp_tac 1); |
|
5670 | 68 |
qed_spec_mp"exec_frag_abstraction"; |
4559 | 69 |
|
70 |
||
6161 | 71 |
Goal "is_abstraction h C A ==> weakeningIOA A C h"; |
4559 | 72 |
by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def])1); |
5132 | 73 |
by Auto_tac; |
4559 | 74 |
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1); |
75 |
(* start state *) |
|
76 |
by (rtac conjI 1); |
|
77 |
by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def,cex_abs_def]) 1); |
|
78 |
(* is-execution-fragment *) |
|
5670 | 79 |
by (etac exec_frag_abstraction 1); |
4559 | 80 |
by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1); |
81 |
qed"abs_is_weakening"; |
|
82 |
||
83 |
||
6161 | 84 |
Goal "[|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |] \ |
4559 | 85 |
\ ==> validIOA C P"; |
5132 | 86 |
by (dtac abs_is_weakening 1); |
4559 | 87 |
by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, |
88 |
validIOA_def, temp_strengthening_def])1); |
|
89 |
by (safe_tac set_cs); |
|
90 |
by (pair_tac "ex" 1); |
|
91 |
qed"AbsRuleT1"; |
|
92 |
||
93 |
||
94 |
(* FIX: Nach TLS.ML *) |
|
95 |
||
5068 | 96 |
Goal "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))"; |
4559 | 97 |
by (simp_tac (simpset() addsimps [IMPLIES_def,temp_sat_def, satisfies_def])1); |
98 |
qed"IMPLIES_temp_sat"; |
|
99 |
||
5068 | 100 |
Goal "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))"; |
4559 | 101 |
by (simp_tac (simpset() addsimps [AND_def,temp_sat_def, satisfies_def])1); |
102 |
qed"AND_temp_sat"; |
|
103 |
||
5068 | 104 |
Goal "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))"; |
4559 | 105 |
by (simp_tac (simpset() addsimps [OR_def,temp_sat_def, satisfies_def])1); |
106 |
qed"OR_temp_sat"; |
|
107 |
||
5068 | 108 |
Goal "(ex |== .~ P) = (~ (ex |== P))"; |
4559 | 109 |
by (simp_tac (simpset() addsimps [NOT_def,temp_sat_def, satisfies_def])1); |
110 |
qed"NOT_temp_sat"; |
|
111 |
||
112 |
Addsimps [IMPLIES_temp_sat,AND_temp_sat,OR_temp_sat,NOT_temp_sat]; |
|
113 |
||
114 |
||
5068 | 115 |
Goalw [is_live_abstraction_def] |
6161 | 116 |
"[|is_live_abstraction h (C,L) (A,M); \ |
4559 | 117 |
\ validLIOA (A,M) Q; temp_strengthening Q P h |] \ |
118 |
\ ==> validLIOA (C,L) P"; |
|
5132 | 119 |
by Auto_tac; |
120 |
by (dtac abs_is_weakening 1); |
|
4559 | 121 |
by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, temp_weakening_def2, |
122 |
validLIOA_def, validIOA_def, temp_strengthening_def])1); |
|
123 |
by (safe_tac set_cs); |
|
124 |
by (pair_tac "ex" 1); |
|
125 |
qed"AbsRuleT2"; |
|
126 |
||
127 |
||
5068 | 128 |
Goalw [is_live_abstraction_def] |
6161 | 129 |
"[|is_live_abstraction h (C,L) (A,M); \ |
4559 | 130 |
\ validLIOA (A,M) (H1 .--> Q); temp_strengthening Q P h; \ |
131 |
\ temp_weakening H1 H2 h; validLIOA (C,L) H2 |] \ |
|
132 |
\ ==> validLIOA (C,L) P"; |
|
5132 | 133 |
by Auto_tac; |
134 |
by (dtac abs_is_weakening 1); |
|
4559 | 135 |
by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, temp_weakening_def2, |
136 |
validLIOA_def, validIOA_def, temp_strengthening_def])1); |
|
137 |
by (safe_tac set_cs); |
|
138 |
by (pair_tac "ex" 1); |
|
139 |
qed"AbsRuleTImprove"; |
|
140 |
||
141 |
||
142 |
section "Correctness of safe abstraction"; |
|
143 |
||
144 |
(* ---------------------------------------------------------------- *) |
|
145 |
(* Correctness of safe abstraction *) |
|
146 |
(* ---------------------------------------------------------------- *) |
|
147 |
||
148 |
||
5068 | 149 |
Goalw [is_abstraction_def,is_ref_map_def] |
6161 | 150 |
"is_abstraction h C A ==> is_ref_map h C A"; |
4559 | 151 |
by (safe_tac set_cs); |
152 |
by (res_inst_tac[("x","(a,h t)>>nil")] exI 1); |
|
153 |
by (asm_full_simp_tac (simpset() addsimps [move_def])1); |
|
154 |
qed"abstraction_is_ref_map"; |
|
155 |
||
156 |
||
6161 | 157 |
Goal "[| inp(C)=inp(A); out(C)=out(A); \ |
4559 | 158 |
\ is_abstraction h C A |] \ |
159 |
\ ==> C =<| A"; |
|
160 |
by (asm_full_simp_tac (simpset() addsimps [ioa_implements_def]) 1); |
|
5132 | 161 |
by (rtac trace_inclusion 1); |
4559 | 162 |
by (simp_tac (simpset() addsimps [externals_def])1); |
163 |
by (SELECT_GOAL (auto_tac (claset(),simpset()))1); |
|
5132 | 164 |
by (etac abstraction_is_ref_map 1); |
4559 | 165 |
qed"abs_safety"; |
166 |
||
167 |
||
168 |
section "Correctness of life abstraction"; |
|
169 |
||
170 |
(* ---------------------------------------------------------------- *) |
|
171 |
(* Correctness of life abstraction *) |
|
172 |
(* ---------------------------------------------------------------- *) |
|
173 |
||
174 |
||
175 |
(* Reduces to Filter (Map fst x) = Filter (Map fst (Map (%(a,t). (a,x)) x), |
|
176 |
that is to special Map Lemma *) |
|
5068 | 177 |
Goalw [cex_abs_def,mk_trace_def,filter_act_def] |
6161 | 178 |
"ext C = ext A \ |
10835 | 179 |
\ ==> mk_trace C$xs = mk_trace A$(snd (cex_abs f (s,xs)))"; |
4559 | 180 |
by (Asm_full_simp_tac 1); |
181 |
by (pair_induct_tac "xs" [] 1); |
|
182 |
qed"traces_coincide_abs"; |
|
183 |
||
184 |
||
185 |
(* Does not work with abstraction_is_ref_map as proof of abs_safety, because |
|
186 |
is_live_abstraction includes temp_strengthening which is necessarily based |
|
187 |
on cex_abs and not on corresp_ex. Thus, the proof is redoone in a more specific |
|
188 |
way for cex_abs *) |
|
6161 | 189 |
Goal "[| inp(C)=inp(A); out(C)=out(A); \ |
4559 | 190 |
\ is_live_abstraction h (C,M) (A,L) |] \ |
191 |
\ ==> live_implements (C,M) (A,L)"; |
|
192 |
||
193 |
by (asm_full_simp_tac (simpset() addsimps [is_live_abstraction_def, live_implements_def, |
|
194 |
livetraces_def,liveexecutions_def]) 1); |
|
195 |
by (safe_tac set_cs); |
|
196 |
by (res_inst_tac[("x","cex_abs h ex")] exI 1); |
|
197 |
by (safe_tac set_cs); |
|
198 |
(* Traces coincide *) |
|
199 |
by (pair_tac "ex" 1); |
|
200 |
by (rtac traces_coincide_abs 1); |
|
201 |
by (simp_tac (simpset() addsimps [externals_def])1); |
|
202 |
by (SELECT_GOAL (auto_tac (claset(),simpset()))1); |
|
203 |
||
204 |
(* cex_abs is execution *) |
|
205 |
by (pair_tac "ex" 1); |
|
206 |
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1); |
|
207 |
(* start state *) |
|
208 |
by (rtac conjI 1); |
|
209 |
by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def,cex_abs_def]) 1); |
|
210 |
(* is-execution-fragment *) |
|
5670 | 211 |
by (etac exec_frag_abstraction 1); |
4559 | 212 |
by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1); |
213 |
||
214 |
(* Liveness *) |
|
215 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2]) 1); |
|
216 |
by (pair_tac "ex" 1); |
|
217 |
qed"abs_liveness"; |
|
218 |
||
219 |
(* FIX: NAch Traces.ML bringen *) |
|
220 |
||
5068 | 221 |
Goalw [ioa_implements_def] |
6161 | 222 |
"[| A =<| B; B =<| C|] ==> A =<| C"; |
5132 | 223 |
by Auto_tac; |
4559 | 224 |
qed"implements_trans"; |
225 |
||
226 |
||
227 |
section "Abstraction Rules for Automata"; |
|
228 |
||
229 |
(* ---------------------------------------------------------------- *) |
|
230 |
(* Abstraction Rules for Automata *) |
|
231 |
(* ---------------------------------------------------------------- *) |
|
232 |
||
6161 | 233 |
Goal "[| inp(C)=inp(A); out(C)=out(A); \ |
4559 | 234 |
\ inp(Q)=inp(P); out(Q)=out(P); \ |
235 |
\ is_abstraction h1 C A; \ |
|
236 |
\ A =<| Q ; \ |
|
237 |
\ is_abstraction h2 Q P |] \ |
|
238 |
\ ==> C =<| P"; |
|
5132 | 239 |
by (dtac abs_safety 1); |
4559 | 240 |
by (REPEAT (atac 1)); |
5132 | 241 |
by (dtac abs_safety 1); |
4559 | 242 |
by (REPEAT (atac 1)); |
5132 | 243 |
by (etac implements_trans 1); |
244 |
by (etac implements_trans 1); |
|
245 |
by (assume_tac 1); |
|
4559 | 246 |
qed"AbsRuleA1"; |
247 |
||
248 |
||
13388 | 249 |
Goal "!!LC. [| inp(C)=inp(A); out(C)=out(A); \ |
4559 | 250 |
\ inp(Q)=inp(P); out(Q)=out(P); \ |
251 |
\ is_live_abstraction h1 (C,LC) (A,LA); \ |
|
252 |
\ live_implements (A,LA) (Q,LQ) ; \ |
|
253 |
\ is_live_abstraction h2 (Q,LQ) (P,LP) |] \ |
|
254 |
\ ==> live_implements (C,LC) (P,LP)"; |
|
5132 | 255 |
by (dtac abs_liveness 1); |
4559 | 256 |
by (REPEAT (atac 1)); |
5132 | 257 |
by (dtac abs_liveness 1); |
4559 | 258 |
by (REPEAT (atac 1)); |
5132 | 259 |
by (etac live_implements_trans 1); |
260 |
by (etac live_implements_trans 1); |
|
261 |
by (assume_tac 1); |
|
4559 | 262 |
qed"AbsRuleA2"; |
263 |
||
264 |
||
265 |
Delsimps [split_paired_All]; |
|
266 |
||
267 |
||
268 |
section "Localizing Temporal Strengthenings and Weakenings"; |
|
269 |
||
270 |
(* ---------------------------------------------------------------- *) |
|
271 |
(* Localizing Temproal Strengthenings - 1 *) |
|
272 |
(* ---------------------------------------------------------------- *) |
|
273 |
||
5068 | 274 |
Goalw [temp_strengthening_def] |
6161 | 275 |
"[| temp_strengthening P1 Q1 h; \ |
4559 | 276 |
\ temp_strengthening P2 Q2 h |] \ |
277 |
\ ==> temp_strengthening (P1 .& P2) (Q1 .& Q2) h"; |
|
5132 | 278 |
by Auto_tac; |
4559 | 279 |
qed"strength_AND"; |
280 |
||
5068 | 281 |
Goalw [temp_strengthening_def] |
6161 | 282 |
"[| temp_strengthening P1 Q1 h; \ |
4559 | 283 |
\ temp_strengthening P2 Q2 h |] \ |
284 |
\ ==> temp_strengthening (P1 .| P2) (Q1 .| Q2) h"; |
|
5132 | 285 |
by Auto_tac; |
4559 | 286 |
qed"strength_OR"; |
287 |
||
5068 | 288 |
Goalw [temp_strengthening_def] |
6161 | 289 |
"[| temp_weakening P Q h |] \ |
4559 | 290 |
\ ==> temp_strengthening (.~ P) (.~ Q) h"; |
291 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
5132 | 292 |
by Auto_tac; |
4559 | 293 |
qed"strength_NOT"; |
294 |
||
5068 | 295 |
Goalw [temp_strengthening_def] |
6161 | 296 |
"[| temp_weakening P1 Q1 h; \ |
4559 | 297 |
\ temp_strengthening P2 Q2 h |] \ |
298 |
\ ==> temp_strengthening (P1 .--> P2) (Q1 .--> Q2) h"; |
|
299 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
300 |
qed"strength_IMPLIES"; |
|
301 |
||
302 |
||
303 |
||
304 |
(* ---------------------------------------------------------------- *) |
|
305 |
(* Localizing Temproal Weakenings - Part 1 *) |
|
306 |
(* ---------------------------------------------------------------- *) |
|
307 |
||
5068 | 308 |
Goal |
6161 | 309 |
"[| temp_weakening P1 Q1 h; \ |
4559 | 310 |
\ temp_weakening P2 Q2 h |] \ |
311 |
\ ==> temp_weakening (P1 .& P2) (Q1 .& Q2) h"; |
|
312 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
313 |
qed"weak_AND"; |
|
314 |
||
5068 | 315 |
Goal |
6161 | 316 |
"[| temp_weakening P1 Q1 h; \ |
4559 | 317 |
\ temp_weakening P2 Q2 h |] \ |
318 |
\ ==> temp_weakening (P1 .| P2) (Q1 .| Q2) h"; |
|
319 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
320 |
qed"weak_OR"; |
|
321 |
||
5068 | 322 |
Goalw [temp_strengthening_def] |
6161 | 323 |
"[| temp_strengthening P Q h |] \ |
4559 | 324 |
\ ==> temp_weakening (.~ P) (.~ Q) h"; |
325 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
5132 | 326 |
by Auto_tac; |
4559 | 327 |
qed"weak_NOT"; |
328 |
||
5068 | 329 |
Goalw [temp_strengthening_def] |
6161 | 330 |
"[| temp_strengthening P1 Q1 h; \ |
4559 | 331 |
\ temp_weakening P2 Q2 h |] \ |
332 |
\ ==> temp_weakening (P1 .--> P2) (Q1 .--> Q2) h"; |
|
333 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1); |
|
334 |
qed"weak_IMPLIES"; |
|
335 |
||
336 |
||
337 |
(* ---------------------------------------------------------------- *) |
|
338 |
(* Localizing Temproal Strengthenings - 2 *) |
|
339 |
(* ---------------------------------------------------------------- *) |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
340 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
341 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
342 |
(* ------------------ Box ----------------------------*) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
343 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
344 |
(* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *) |
5068 | 345 |
Goal "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
346 |
by (Seq_case_simp_tac "x" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
347 |
by Auto_tac; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
348 |
qed"UU_is_Conc"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
349 |
|
5068 | 350 |
Goal |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
351 |
"Finite s1 --> \ |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
352 |
\ (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
353 |
by (rtac impI 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
354 |
by (Seq_Finite_induct_tac 1); |
5670 | 355 |
by (Blast_tac 1); |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
356 |
(* main case *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
357 |
by (clarify_tac set_cs 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
358 |
by (pair_tac "ex" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
359 |
by (Seq_case_simp_tac "y" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
360 |
(* UU case *) |
5676 | 361 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc])1); |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
362 |
(* nil case *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
363 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc])1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
364 |
(* cons case *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
365 |
by (pair_tac "aa" 1); |
5132 | 366 |
by Auto_tac; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
367 |
qed_spec_mp"ex2seqConc"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
368 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
369 |
(* important property of ex2seq: can be shiftet, as defined "pointwise" *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
370 |
|
5068 | 371 |
Goalw [tsuffix_def,suffix_def] |
6161 | 372 |
"tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')"; |
5132 | 373 |
by Auto_tac; |
374 |
by (dtac ex2seqConc 1); |
|
375 |
by Auto_tac; |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
376 |
qed"ex2seq_tsuffix"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
377 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
378 |
|
5976 | 379 |
(* FIX: NAch Sequence.ML bringen *) |
380 |
||
10835 | 381 |
Goal "(Map f$s = nil) = (s=nil)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
382 |
by (Seq_case_simp_tac "s" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
383 |
qed"Mapnil"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
384 |
|
10835 | 385 |
Goal "(Map f$s = UU) = (s=UU)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
386 |
by (Seq_case_simp_tac "s" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
387 |
qed"MapUU"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
388 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
389 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
390 |
(* important property of cex_absSeq: As it is a 1to1 correspondence, |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
391 |
properties carry over *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
392 |
|
5068 | 393 |
Goalw [tsuffix_def,suffix_def,cex_absSeq_def] |
6161 | 394 |
"tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)"; |
5132 | 395 |
by Auto_tac; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
396 |
by (asm_full_simp_tac (simpset() addsimps [Mapnil])1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
397 |
by (asm_full_simp_tac (simpset() addsimps [MapUU])1); |
10835 | 398 |
by (res_inst_tac [("x","Map (%(s,a,t). (h s,a, h t))$s1")] exI 1); |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
399 |
by (asm_full_simp_tac (simpset() addsimps [Map2Finite,MapConc])1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
400 |
qed"cex_absSeq_tsuffix"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
401 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
402 |
|
5068 | 403 |
Goalw [temp_strengthening_def,state_strengthening_def, temp_sat_def, |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
404 |
satisfies_def,Box_def] |
6161 | 405 |
"[| temp_strengthening P Q h |]\ |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
406 |
\ ==> temp_strengthening ([] P) ([] Q) h"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
407 |
by (clarify_tac set_cs 1); |
7499 | 408 |
by (ftac ex2seq_tsuffix 1); |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
409 |
by (clarify_tac set_cs 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
410 |
by (dres_inst_tac [("h","h")] cex_absSeq_tsuffix 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
411 |
by (asm_full_simp_tac (simpset() addsimps [ex2seq_abs_cex])1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
412 |
qed"strength_Box"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
413 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
414 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
415 |
(* ------------------ Init ----------------------------*) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
416 |
|
5068 | 417 |
Goalw [temp_strengthening_def,state_strengthening_def, |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
418 |
temp_sat_def,satisfies_def,Init_def,unlift_def] |
6161 | 419 |
"[| state_strengthening P Q h |]\ |
4559 | 420 |
\ ==> temp_strengthening (Init P) (Init Q) h"; |
421 |
by (safe_tac set_cs); |
|
422 |
by (pair_tac "ex" 1); |
|
423 |
by (Seq_case_simp_tac "y" 1); |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
424 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
425 |
qed"strength_Init"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
426 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
427 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
428 |
(* ------------------ Next ----------------------------*) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
429 |
|
5068 | 430 |
Goal |
10835 | 431 |
"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
432 |
by (pair_tac "ex" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
433 |
by (Seq_case_simp_tac "y" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
434 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
435 |
by (Seq_case_simp_tac "s" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
436 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
437 |
qed"TL_ex2seq_UU"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
438 |
|
5068 | 439 |
Goal |
10835 | 440 |
"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
441 |
by (pair_tac "ex" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
442 |
by (Seq_case_simp_tac "y" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
443 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
444 |
by (Seq_case_simp_tac "s" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
445 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
446 |
qed"TL_ex2seq_nil"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
447 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
448 |
(* FIX: put to Sequence Lemmas *) |
10835 | 449 |
Goal "Map f$(TL$s) = TL$(Map f$s)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
450 |
by (Seq_induct_tac "s" [] 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
451 |
qed"MapTL"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
452 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
453 |
(* important property of cex_absSeq: As it is a 1to1 correspondence, |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
454 |
properties carry over *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
455 |
|
5068 | 456 |
Goalw [cex_absSeq_def] |
10835 | 457 |
"cex_absSeq h (TL$s) = (TL$(cex_absSeq h s))"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
458 |
by (simp_tac (simpset() addsimps [MapTL]) 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
459 |
qed"cex_absSeq_TL"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
460 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
461 |
(* important property of ex2seq: can be shiftet, as defined "pointwise" *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
462 |
|
10835 | 463 |
Goal "[| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL$(ex2seq ex) = ex2seq ex')"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
464 |
by (pair_tac "ex" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
465 |
by (Seq_case_simp_tac "y" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
466 |
by (pair_tac "a" 1); |
5132 | 467 |
by Auto_tac; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
468 |
qed"TLex2seq"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
469 |
|
5670 | 470 |
|
10835 | 471 |
Goal "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
472 |
by (pair_tac "ex" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
473 |
by (Seq_case_simp_tac "y" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
474 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
475 |
by (Seq_case_simp_tac "s" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
476 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
477 |
qed"ex2seqnilTL"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
478 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
479 |
|
5068 | 480 |
Goalw [temp_strengthening_def,state_strengthening_def, |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
481 |
temp_sat_def, satisfies_def,Next_def] |
6161 | 482 |
"[| temp_strengthening P Q h |]\ |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
483 |
\ ==> temp_strengthening (Next P) (Next Q) h"; |
4833 | 484 |
by (Asm_full_simp_tac 1); |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
485 |
by (safe_tac set_cs); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
486 |
by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
487 |
by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
488 |
by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
489 |
by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
490 |
(* cons case *) |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
491 |
by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU, |
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
5676
diff
changeset
|
492 |
ex2seq_abs_cex,cex_absSeq_TL RS sym, ex2seqnilTL])1); |
6161 | 493 |
by (etac conjE 1); |
5132 | 494 |
by (dtac TLex2seq 1); |
495 |
by (assume_tac 1); |
|
496 |
by Auto_tac; |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
497 |
qed"strength_Next"; |
4559 | 498 |
|
499 |
||
500 |
||
501 |
(* ---------------------------------------------------------------- *) |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
502 |
(* Localizing Temporal Weakenings - 2 *) |
4559 | 503 |
(* ---------------------------------------------------------------- *) |
504 |
||
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
505 |
|
5068 | 506 |
Goal |
6161 | 507 |
"[| state_weakening P Q h |]\ |
4559 | 508 |
\ ==> temp_weakening (Init P) (Init Q) h"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
509 |
by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2, |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
510 |
state_weakening_def2, temp_sat_def,satisfies_def,Init_def,unlift_def])1); |
4559 | 511 |
by (safe_tac set_cs); |
512 |
by (pair_tac "ex" 1); |
|
513 |
by (Seq_case_simp_tac "y" 1); |
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
514 |
by (pair_tac "a" 1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
515 |
qed"weak_Init"; |
4559 | 516 |
|
517 |
||
518 |
(* ---------------------------------------------------------------- *) |
|
519 |
(* Localizing Temproal Strengthenings - 3 *) |
|
520 |
(* ---------------------------------------------------------------- *) |
|
521 |
||
522 |
||
5068 | 523 |
Goalw [Diamond_def] |
6161 | 524 |
"[| temp_strengthening P Q h |]\ |
4559 | 525 |
\ ==> temp_strengthening (<> P) (<> Q) h"; |
5132 | 526 |
by (rtac strength_NOT 1); |
527 |
by (rtac weak_Box 1); |
|
528 |
by (etac weak_NOT 1); |
|
4559 | 529 |
qed"strength_Diamond"; |
530 |
||
5068 | 531 |
Goalw [Leadsto_def] |
6161 | 532 |
"[| temp_weakening P1 P2 h;\ |
4559 | 533 |
\ temp_strengthening Q1 Q2 h |]\ |
534 |
\ ==> temp_strengthening (P1 ~> Q1) (P2 ~> Q2) h"; |
|
5132 | 535 |
by (rtac strength_Box 1); |
536 |
by (etac strength_IMPLIES 1); |
|
537 |
by (etac strength_Diamond 1); |
|
4559 | 538 |
qed"strength_Leadsto"; |
539 |
||
540 |
||
541 |
(* ---------------------------------------------------------------- *) |
|
542 |
(* Localizing Temporal Weakenings - 3 *) |
|
543 |
(* ---------------------------------------------------------------- *) |
|
544 |
||
545 |
||
5068 | 546 |
Goalw [Diamond_def] |
6161 | 547 |
"[| temp_weakening P Q h |]\ |
4559 | 548 |
\ ==> temp_weakening (<> P) (<> Q) h"; |
5132 | 549 |
by (rtac weak_NOT 1); |
550 |
by (rtac strength_Box 1); |
|
551 |
by (etac strength_NOT 1); |
|
4559 | 552 |
qed"weak_Diamond"; |
553 |
||
5068 | 554 |
Goalw [Leadsto_def] |
6161 | 555 |
"[| temp_strengthening P1 P2 h;\ |
4559 | 556 |
\ temp_weakening Q1 Q2 h |]\ |
557 |
\ ==> temp_weakening (P1 ~> Q1) (P2 ~> Q2) h"; |
|
5132 | 558 |
by (rtac weak_Box 1); |
559 |
by (etac weak_IMPLIES 1); |
|
560 |
by (etac weak_Diamond 1); |
|
4559 | 561 |
qed"weak_Leadsto"; |
562 |
||
5068 | 563 |
Goalw [WF_def] |
4559 | 564 |
" !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \ |
565 |
\ ==> temp_weakening (WF A acts) (WF C acts) h"; |
|
5132 | 566 |
by (rtac weak_IMPLIES 1); |
567 |
by (rtac strength_Diamond 1); |
|
568 |
by (rtac strength_Box 1); |
|
569 |
by (rtac strength_Init 1); |
|
570 |
by (rtac weak_Box 2); |
|
571 |
by (rtac weak_Diamond 2); |
|
572 |
by (rtac weak_Init 2); |
|
4559 | 573 |
by (auto_tac (claset(), |
574 |
simpset() addsimps [state_weakening_def,state_strengthening_def, |
|
575 |
xt2_def,plift_def,option_lift_def,NOT_def])); |
|
576 |
qed"weak_WF"; |
|
577 |
||
5068 | 578 |
Goalw [SF_def] |
4559 | 579 |
" !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \ |
580 |
\ ==> temp_weakening (SF A acts) (SF C acts) h"; |
|
5132 | 581 |
by (rtac weak_IMPLIES 1); |
582 |
by (rtac strength_Box 1); |
|
583 |
by (rtac strength_Diamond 1); |
|
584 |
by (rtac strength_Init 1); |
|
585 |
by (rtac weak_Box 2); |
|
586 |
by (rtac weak_Diamond 2); |
|
587 |
by (rtac weak_Init 2); |
|
4559 | 588 |
by (auto_tac (claset(), |
589 |
simpset() addsimps [state_weakening_def,state_strengthening_def, |
|
590 |
xt2_def,plift_def,option_lift_def,NOT_def])); |
|
591 |
qed"weak_SF"; |
|
592 |
||
593 |
||
594 |
val weak_strength_lemmas = |
|
595 |
[weak_OR,weak_AND,weak_NOT,weak_IMPLIES,weak_Box,weak_Next,weak_Init, |
|
596 |
weak_Diamond,weak_Leadsto,strength_OR,strength_AND,strength_NOT, |
|
597 |
strength_IMPLIES,strength_Box,strength_Next,strength_Init, |
|
598 |
strength_Diamond,strength_Leadsto,weak_WF,weak_SF]; |
|
599 |
||
600 |
fun abstraction_tac i = |
|
601 |
SELECT_GOAL (auto_tac (claset() addSIs weak_strength_lemmas, |
|
4725 | 602 |
simpset() addsimps [state_strengthening_def,state_weakening_def])) i; |