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