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