author | wenzelm |
Tue, 29 Aug 2000 20:13:45 +0200 | |
changeset 9731 | 3eb72671e5db |
parent 6161 | bc2a76ce1ea3 |
child 11655 | 923e4d0d36d5 |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/Automata.ML |
3275 | 2 |
ID: $Id$ |
3071 | 3 |
Author: Olaf Mueller, Tobias Nipkow, Konrad Slind |
4 |
Copyright 1994, 1996 TU Muenchen |
|
5 |
||
6 |
The I/O automata of Lynch and Tuttle. |
|
7 |
*) |
|
8 |
||
4814 | 9 |
(* this modification of the simpset is local to this file *) |
4536 | 10 |
Delsimps [split_paired_Ex]; |
3071 | 11 |
|
3521 | 12 |
val ioa_projections = [asig_of_def, starts_of_def, trans_of_def,wfair_of_def,sfair_of_def]; |
3071 | 13 |
|
6161 | 14 |
(* ------------------------------------------------------------------------- *) |
3071 | 15 |
|
16 |
section "asig_of, starts_of, trans_of"; |
|
17 |
||
18 |
||
5068 | 19 |
Goal |
3521 | 20 |
"((asig_of (x,y,z,w,s)) = x) & \ |
21 |
\ ((starts_of (x,y,z,w,s)) = y) & \ |
|
22 |
\ ((trans_of (x,y,z,w,s)) = z) & \ |
|
23 |
\ ((wfair_of (x,y,z,w,s)) = w) & \ |
|
24 |
\ ((sfair_of (x,y,z,w,s)) = s)"; |
|
4098 | 25 |
by (simp_tac (simpset() addsimps ioa_projections) 1); |
3071 | 26 |
qed "ioa_triple_proj"; |
27 |
||
5068 | 28 |
Goalw [is_trans_of_def,actions_def, is_asig_def] |
6161 | 29 |
"[| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A"; |
3071 | 30 |
by (REPEAT(etac conjE 1)); |
31 |
by (EVERY1[etac allE, etac impE, atac]); |
|
32 |
by (Asm_full_simp_tac 1); |
|
33 |
qed "trans_in_actions"; |
|
34 |
||
5068 | 35 |
Goal |
3071 | 36 |
"starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"; |
4098 | 37 |
by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1); |
3071 | 38 |
qed "starts_of_par"; |
39 |
||
5068 | 40 |
Goal |
3521 | 41 |
"trans_of(A || B) = {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) \ |
42 |
\ in (a:act A | a:act B) & \ |
|
43 |
\ (if a:act A then \ |
|
44 |
\ (fst(s),a,fst(t)):trans_of(A) \ |
|
45 |
\ else fst(t) = fst(s)) \ |
|
46 |
\ & \ |
|
47 |
\ (if a:act B then \ |
|
48 |
\ (snd(s),a,snd(t)):trans_of(B) \ |
|
49 |
\ else snd(t) = snd(s))}"; |
|
50 |
||
4423 | 51 |
by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1); |
3521 | 52 |
qed "trans_of_par"; |
53 |
||
3071 | 54 |
|
6161 | 55 |
(* ------------------------------------------------------------------------- *) |
3071 | 56 |
|
57 |
section "actions and par"; |
|
58 |
||
59 |
||
5068 | 60 |
Goal |
3071 | 61 |
"actions(asig_comp a b) = actions(a) Un actions(b)"; |
4098 | 62 |
by (simp_tac (simpset() addsimps |
3071 | 63 |
([actions_def,asig_comp_def]@asig_projections)) 1); |
64 |
by (fast_tac (set_cs addSIs [equalityI]) 1); |
|
65 |
qed "actions_asig_comp"; |
|
66 |
||
67 |
||
5068 | 68 |
Goal "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)"; |
4098 | 69 |
by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1); |
3071 | 70 |
qed "asig_of_par"; |
71 |
||
72 |
||
5068 | 73 |
Goal "ext (A1||A2) = \ |
3071 | 74 |
\ (ext A1) Un (ext A2)"; |
4098 | 75 |
by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_of_par,asig_comp_def, |
3071 | 76 |
asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1); |
77 |
by (rtac set_ext 1); |
|
78 |
by (fast_tac set_cs 1); |
|
79 |
qed"externals_of_par"; |
|
80 |
||
5068 | 81 |
Goal "act (A1||A2) = \ |
3071 | 82 |
\ (act A1) Un (act A2)"; |
4098 | 83 |
by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def, |
3071 | 84 |
asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1); |
85 |
by (rtac set_ext 1); |
|
86 |
by (fast_tac set_cs 1); |
|
87 |
qed"actions_of_par"; |
|
88 |
||
5068 | 89 |
Goal "inp (A1||A2) =\ |
3071 | 90 |
\ ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))"; |
4098 | 91 |
by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def, |
3071 | 92 |
asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1); |
93 |
qed"inputs_of_par"; |
|
94 |
||
5068 | 95 |
Goal "out (A1||A2) =\ |
3071 | 96 |
\ (out A1) Un (out A2)"; |
4098 | 97 |
by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def, |
3071 | 98 |
asig_outputs_def,Un_def,set_diff_def]) 1); |
99 |
qed"outputs_of_par"; |
|
100 |
||
5068 | 101 |
Goal "int (A1||A2) =\ |
3656 | 102 |
\ (int A1) Un (int A2)"; |
4098 | 103 |
by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def, |
3656 | 104 |
asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1); |
105 |
qed"internals_of_par"; |
|
3071 | 106 |
|
6161 | 107 |
(* ------------------------------------------------------------------------ *) |
3071 | 108 |
|
109 |
section "actions and compatibility"; |
|
110 |
||
5068 | 111 |
Goal"compatible A B = compatible B A"; |
4098 | 112 |
by (asm_full_simp_tac (simpset() addsimps [compatible_def,Int_commute]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
113 |
by Auto_tac; |
3071 | 114 |
qed"compat_commute"; |
115 |
||
5068 | 116 |
Goalw [externals_def,actions_def,compatible_def] |
6161 | 117 |
"[| compatible A1 A2; a:ext A1|] ==> a~:int A2"; |
3071 | 118 |
by (Asm_full_simp_tac 1); |
6161 | 119 |
by (Blast_tac 1); |
3071 | 120 |
qed"ext1_is_not_int2"; |
121 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
122 |
(* just commuting the previous one: better commute compatible *) |
5068 | 123 |
Goalw [externals_def,actions_def,compatible_def] |
6161 | 124 |
"[| compatible A2 A1 ; a:ext A1|] ==> a~:int A2"; |
3071 | 125 |
by (Asm_full_simp_tac 1); |
6161 | 126 |
by (Blast_tac 1); |
3071 | 127 |
qed"ext2_is_not_int1"; |
128 |
||
129 |
bind_thm("ext1_ext2_is_not_act2",ext1_is_not_int2 RS int_and_ext_is_act); |
|
130 |
bind_thm("ext1_ext2_is_not_act1",ext2_is_not_int1 RS int_and_ext_is_act); |
|
131 |
||
5068 | 132 |
Goalw [externals_def,actions_def,compatible_def] |
6161 | 133 |
"[| compatible A B; x:int A |] ==> x~:ext B"; |
3071 | 134 |
by (Asm_full_simp_tac 1); |
6161 | 135 |
by (Blast_tac 1); |
3071 | 136 |
qed"intA_is_not_extB"; |
137 |
||
5068 | 138 |
Goalw [externals_def,actions_def,compatible_def,is_asig_def,asig_of_def] |
6161 | 139 |
"[| compatible A B; a:int A |] ==> a ~: act B"; |
3071 | 140 |
by (Asm_full_simp_tac 1); |
6161 | 141 |
by (Blast_tac 1); |
3071 | 142 |
qed"intA_is_not_actB"; |
143 |
||
3521 | 144 |
(* the only one that needs disjointness of outputs and of internals and _all_ acts *) |
5068 | 145 |
Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def, |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
146 |
compatible_def,is_asig_def,asig_of_def] |
6161 | 147 |
"[| compatible A B; a:out A ;a:act B|] ==> a : inp B"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
148 |
by (Asm_full_simp_tac 1); |
6161 | 149 |
by (Blast_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
150 |
qed"outAactB_is_inpB"; |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
151 |
|
3521 | 152 |
(* needed for propagation of input_enabledness from A,B to A||B *) |
5068 | 153 |
Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def, |
3521 | 154 |
compatible_def,is_asig_def,asig_of_def] |
6161 | 155 |
"[| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B"; |
3521 | 156 |
by (Asm_full_simp_tac 1); |
6161 | 157 |
by (Blast_tac 1); |
3521 | 158 |
qed"inpAAactB_is_inpBoroutB"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
159 |
|
6161 | 160 |
(* ------------------------------------------------------------------------- *) |
3521 | 161 |
|
162 |
section "input_enabledness and par"; |
|
163 |
||
164 |
||
165 |
(* ugly case distinctions. Heart of proof: |
|
166 |
1. inpAAactB_is_inpBoroutB ie. internals are really hidden. |
|
167 |
2. inputs_of_par: outputs are no longer inputs of par. This is important here *) |
|
5068 | 168 |
Goalw [input_enabled_def] |
6161 | 169 |
"[| compatible A B; input_enabled A; input_enabled B|] \ |
3521 | 170 |
\ ==> input_enabled (A||B)"; |
4814 | 171 |
by (asm_full_simp_tac (simpset()addsimps[Let_def,inputs_of_par,trans_of_par])1); |
3521 | 172 |
by (safe_tac set_cs); |
4098 | 173 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
174 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 2); |
|
3521 | 175 |
(* a: inp A *) |
176 |
by (case_tac "a:act B" 1); |
|
177 |
(* a:act B *) |
|
178 |
by (eres_inst_tac [("x","a")] allE 1); |
|
179 |
by (Asm_full_simp_tac 1); |
|
4423 | 180 |
by (dtac inpAAactB_is_inpBoroutB 1); |
181 |
by (assume_tac 1); |
|
182 |
by (assume_tac 1); |
|
3521 | 183 |
by (eres_inst_tac [("x","a")] allE 1); |
184 |
by (Asm_full_simp_tac 1); |
|
185 |
by (eres_inst_tac [("x","aa")] allE 1); |
|
186 |
by (eres_inst_tac [("x","b")] allE 1); |
|
4423 | 187 |
by (etac exE 1); |
188 |
by (etac exE 1); |
|
3521 | 189 |
by (res_inst_tac [("x","(s2,s2a)")] exI 1); |
4098 | 190 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
3521 | 191 |
(* a~: act B*) |
4098 | 192 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
3521 | 193 |
by (eres_inst_tac [("x","a")] allE 1); |
194 |
by (Asm_full_simp_tac 1); |
|
195 |
by (eres_inst_tac [("x","aa")] allE 1); |
|
4423 | 196 |
by (etac exE 1); |
3521 | 197 |
by (res_inst_tac [("x","(s2,b)")] exI 1); |
198 |
by (Asm_full_simp_tac 1); |
|
199 |
||
200 |
(* a:inp B *) |
|
201 |
by (case_tac "a:act A" 1); |
|
202 |
(* a:act A *) |
|
203 |
by (eres_inst_tac [("x","a")] allE 1); |
|
204 |
by (eres_inst_tac [("x","a")] allE 1); |
|
4098 | 205 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
4473 | 206 |
by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1); |
4423 | 207 |
by (dtac inpAAactB_is_inpBoroutB 1); |
3521 | 208 |
back(); |
4423 | 209 |
by (assume_tac 1); |
210 |
by (assume_tac 1); |
|
3521 | 211 |
by (Asm_full_simp_tac 1); |
212 |
by (rotate_tac ~1 1); |
|
213 |
by (Asm_full_simp_tac 1); |
|
214 |
by (eres_inst_tac [("x","aa")] allE 1); |
|
215 |
by (eres_inst_tac [("x","b")] allE 1); |
|
4423 | 216 |
by (etac exE 1); |
217 |
by (etac exE 1); |
|
3521 | 218 |
by (res_inst_tac [("x","(s2,s2a)")] exI 1); |
4098 | 219 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
3521 | 220 |
(* a~: act B*) |
4098 | 221 |
by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1); |
3521 | 222 |
by (eres_inst_tac [("x","a")] allE 1); |
223 |
by (Asm_full_simp_tac 1); |
|
224 |
by (eres_inst_tac [("x","a")] allE 1); |
|
225 |
by (Asm_full_simp_tac 1); |
|
226 |
by (eres_inst_tac [("x","b")] allE 1); |
|
4423 | 227 |
by (etac exE 1); |
3521 | 228 |
by (res_inst_tac [("x","(aa,s2)")] exI 1); |
229 |
by (Asm_full_simp_tac 1); |
|
230 |
qed"input_enabled_par"; |
|
3071 | 231 |
|
6161 | 232 |
(* ------------------------------------------------------------------------- *) |
3071 | 233 |
|
234 |
section "invariants"; |
|
235 |
||
236 |
val [p1,p2] = goalw thy [invariant_def] |
|
237 |
"[| !!s. s:starts_of(A) ==> P(s); \ |
|
238 |
\ !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \ |
|
239 |
\ ==> invariant A P"; |
|
240 |
by (rtac allI 1); |
|
241 |
by (rtac impI 1); |
|
5102 | 242 |
by (res_inst_tac [("xa","s")] reachable.induct 1); |
3071 | 243 |
by (atac 1); |
244 |
by (etac p1 1); |
|
245 |
by (eres_inst_tac [("s1","sa")] (p2 RS mp) 1); |
|
246 |
by (REPEAT (atac 1)); |
|
247 |
qed"invariantI"; |
|
248 |
||
249 |
||
250 |
val [p1,p2] = goal thy |
|
251 |
"[| !!s. s : starts_of(A) ==> P(s); \ |
|
6161 | 252 |
\ !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) \ |
3071 | 253 |
\ |] ==> invariant A P"; |
254 |
by (fast_tac (HOL_cs addSIs [invariantI] addSDs [p1,p2]) 1); |
|
255 |
qed "invariantI1"; |
|
256 |
||
6161 | 257 |
Goalw [invariant_def] "[| invariant A P; reachable A s |] ==> P(s)"; |
258 |
by (Blast_tac 1); |
|
3071 | 259 |
qed "invariantE"; |
260 |
||
6161 | 261 |
(* ------------------------------------------------------------------------- *) |
3071 | 262 |
|
263 |
section "restrict"; |
|
264 |
||
265 |
||
6161 | 266 |
val reachable_0 = reachable.reachable_0 |
267 |
and reachable_n = reachable.reachable_n; |
|
268 |
||
5068 | 269 |
Goal "starts_of(restrict ioa acts) = starts_of(ioa) & \ |
3656 | 270 |
\ trans_of(restrict ioa acts) = trans_of(ioa)"; |
4098 | 271 |
by (simp_tac (simpset() addsimps ([restrict_def]@ioa_projections)) 1); |
3071 | 272 |
qed "cancel_restrict_a"; |
273 |
||
5068 | 274 |
Goal "reachable (restrict ioa acts) s = reachable ioa s"; |
3071 | 275 |
by (rtac iffI 1); |
3457 | 276 |
by (etac reachable.induct 1); |
4098 | 277 |
by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a,reachable_0]) 1); |
3071 | 278 |
by (etac reachable_n 1); |
4098 | 279 |
by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1); |
3071 | 280 |
(* <-- *) |
3457 | 281 |
by (etac reachable.induct 1); |
3071 | 282 |
by (rtac reachable_0 1); |
4098 | 283 |
by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1); |
3071 | 284 |
by (etac reachable_n 1); |
4098 | 285 |
by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1); |
3071 | 286 |
qed "cancel_restrict_b"; |
287 |
||
5068 | 288 |
Goal "act (restrict A acts) = act A"; |
4098 | 289 |
by (simp_tac (simpset() addsimps [actions_def,asig_internals_def, |
3656 | 290 |
asig_outputs_def,asig_inputs_def,externals_def,asig_of_def,restrict_def, |
291 |
restrict_asig_def]) 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
292 |
by Auto_tac; |
3656 | 293 |
qed"acts_restrict"; |
294 |
||
5068 | 295 |
Goal "starts_of(restrict ioa acts) = starts_of(ioa) & \ |
3656 | 296 |
\ trans_of(restrict ioa acts) = trans_of(ioa) & \ |
297 |
\ reachable (restrict ioa acts) s = reachable ioa s & \ |
|
298 |
\ act (restrict A acts) = act A"; |
|
4098 | 299 |
by (simp_tac (simpset() addsimps [cancel_restrict_a,cancel_restrict_b,acts_restrict]) 1); |
3071 | 300 |
qed"cancel_restrict"; |
301 |
||
6161 | 302 |
(* ------------------------------------------------------------------------- *) |
3071 | 303 |
|
304 |
section "rename"; |
|
305 |
||
306 |
||
307 |
||
6161 | 308 |
Goal "s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)"; |
4098 | 309 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,rename_def,trans_of_def]) 1); |
3071 | 310 |
qed"trans_rename"; |
311 |
||
312 |
||
6161 | 313 |
Goal "[| reachable (rename C g) s |] ==> reachable C s"; |
3457 | 314 |
by (etac reachable.induct 1); |
315 |
by (rtac reachable_0 1); |
|
4098 | 316 |
by (asm_full_simp_tac (simpset() addsimps [rename_def]@ioa_projections) 1); |
3457 | 317 |
by (dtac trans_rename 1); |
318 |
by (etac exE 1); |
|
319 |
by (etac conjE 1); |
|
320 |
by (etac reachable_n 1); |
|
321 |
by (assume_tac 1); |
|
3071 | 322 |
qed"reachable_rename"; |
323 |
||
324 |
||
325 |
||
6161 | 326 |
(* ------------------------------------------------------------------------- *) |
3071 | 327 |
|
328 |
section "trans_of(A||B)"; |
|
329 |
||
330 |
||
6161 | 331 |
Goal "[|(s,a,t):trans_of (A||B); a:act A|] \ |
3071 | 332 |
\ ==> (fst s,a,fst t):trans_of A"; |
4098 | 333 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 334 |
qed"trans_A_proj"; |
335 |
||
6161 | 336 |
Goal "[|(s,a,t):trans_of (A||B); a:act B|] \ |
3071 | 337 |
\ ==> (snd s,a,snd t):trans_of B"; |
4098 | 338 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 339 |
qed"trans_B_proj"; |
340 |
||
6161 | 341 |
Goal "[|(s,a,t):trans_of (A||B); a~:act A|]\ |
3071 | 342 |
\ ==> fst s = fst t"; |
4098 | 343 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 344 |
qed"trans_A_proj2"; |
345 |
||
6161 | 346 |
Goal "[|(s,a,t):trans_of (A||B); a~:act B|]\ |
3071 | 347 |
\ ==> snd s = snd t"; |
4098 | 348 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 349 |
qed"trans_B_proj2"; |
350 |
||
6161 | 351 |
Goal "(s,a,t):trans_of (A||B) \ |
3071 | 352 |
\ ==> a :act A | a :act B"; |
4098 | 353 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 354 |
qed"trans_AB_proj"; |
355 |
||
6161 | 356 |
Goal "[|a:act A;a:act B;\ |
3071 | 357 |
\ (fst s,a,fst t):trans_of A;(snd s,a,snd t):trans_of B|]\ |
358 |
\ ==> (s,a,t):trans_of (A||B)"; |
|
4098 | 359 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 360 |
qed"trans_AB"; |
361 |
||
6161 | 362 |
Goal "[|a:act A;a~:act B;\ |
3071 | 363 |
\ (fst s,a,fst t):trans_of A;snd s=snd t|]\ |
364 |
\ ==> (s,a,t):trans_of (A||B)"; |
|
4098 | 365 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 366 |
qed"trans_A_notB"; |
367 |
||
6161 | 368 |
Goal "[|a~:act A;a:act B;\ |
3071 | 369 |
\ (snd s,a,snd t):trans_of B;fst s=fst t|]\ |
370 |
\ ==> (s,a,t):trans_of (A||B)"; |
|
4098 | 371 |
by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1); |
3071 | 372 |
qed"trans_notA_B"; |
373 |
||
374 |
val trans_of_defs1 = [trans_AB,trans_A_notB,trans_notA_B]; |
|
375 |
val trans_of_defs2 = [trans_A_proj,trans_B_proj,trans_A_proj2, |
|
376 |
trans_B_proj2,trans_AB_proj]; |
|
377 |
||
378 |
||
5068 | 379 |
Goal |
3071 | 380 |
"(s,a,t) : trans_of(A || B || C || D) = \ |
381 |
\ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | \ |
|
382 |
\ a:actions(asig_of(D))) & \ |
|
383 |
\ (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) \ |
|
384 |
\ else fst t=fst s) & \ |
|
385 |
\ (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) \ |
|
386 |
\ else fst(snd(t))=fst(snd(s))) & \ |
|
387 |
\ (if a:actions(asig_of(C)) then \ |
|
388 |
\ (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) \ |
|
389 |
\ else fst(snd(snd(t)))=fst(snd(snd(s)))) & \ |
|
390 |
\ (if a:actions(asig_of(D)) then \ |
|
391 |
\ (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) \ |
|
392 |
\ else snd(snd(snd(t)))=snd(snd(snd(s)))))"; |
|
393 |
||
4098 | 394 |
by (simp_tac (simpset() addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq,Let_def]@ |
4833 | 395 |
ioa_projections)) 1); |
3071 | 396 |
qed "trans_of_par4"; |
397 |
||
398 |
||
6161 | 399 |
(* ------------------------------------------------------------------------- *) |
3071 | 400 |
|
3521 | 401 |
section "proof obligation generator for IOA requirements"; |
402 |
||
403 |
(* without assumptions on A and B because is_trans_of is also incorporated in ||def *) |
|
5068 | 404 |
Goalw [is_trans_of_def] "is_trans_of (A||B)"; |
4814 | 405 |
by (simp_tac (simpset() addsimps [Let_def,actions_of_par,trans_of_par]) 1); |
3521 | 406 |
qed"is_trans_of_par"; |
407 |
||
5068 | 408 |
Goalw [is_trans_of_def] |
6161 | 409 |
"is_trans_of A ==> is_trans_of (restrict A acts)"; |
4098 | 410 |
by (asm_simp_tac (simpset() addsimps [cancel_restrict,acts_restrict])1); |
3656 | 411 |
qed"is_trans_of_restrict"; |
3521 | 412 |
|
5068 | 413 |
Goalw [is_trans_of_def,restrict_def,restrict_asig_def] |
6161 | 414 |
"is_trans_of A ==> is_trans_of (rename A f)"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
415 |
by (asm_full_simp_tac |
4814 | 416 |
(simpset() addsimps [Let_def,actions_def,trans_of_def, asig_internals_def, |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
417 |
asig_outputs_def,asig_inputs_def,externals_def, |
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
418 |
asig_of_def,rename_def,rename_set_def]) 1); |
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
419 |
by (Blast_tac 1); |
3656 | 420 |
qed"is_trans_of_rename"; |
3521 | 421 |
|
6161 | 422 |
Goal "[| is_asig_of A; is_asig_of B; compatible A B|] \ |
3521 | 423 |
\ ==> is_asig_of (A||B)"; |
4098 | 424 |
by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def,asig_of_par, |
3656 | 425 |
asig_comp_def,compatible_def,asig_internals_def,asig_outputs_def, |
426 |
asig_inputs_def,actions_def,is_asig_def]) 1); |
|
4098 | 427 |
by (asm_full_simp_tac (simpset() addsimps [asig_of_def]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
428 |
by Auto_tac; |
3656 | 429 |
qed"is_asig_of_par"; |
3521 | 430 |
|
5068 | 431 |
Goalw [is_asig_of_def,is_asig_def,asig_of_def,restrict_def,restrict_asig_def, |
3656 | 432 |
asig_internals_def,asig_outputs_def,asig_inputs_def,externals_def,o_def] |
6161 | 433 |
"is_asig_of A ==> is_asig_of (restrict A f)"; |
3656 | 434 |
by (Asm_full_simp_tac 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
435 |
by Auto_tac; |
3656 | 436 |
qed"is_asig_of_restrict"; |
3521 | 437 |
|
6161 | 438 |
Goal "is_asig_of A ==> is_asig_of (rename A f)"; |
4098 | 439 |
by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def, |
3656 | 440 |
rename_def,rename_set_def,asig_internals_def,asig_outputs_def, |
441 |
asig_inputs_def,actions_def,is_asig_def,asig_of_def]) 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
442 |
by Auto_tac; |
4814 | 443 |
by (ALLGOALS(dres_inst_tac [("s","Some ?x")] sym THEN' Asm_full_simp_tac)); |
6161 | 444 |
by (ALLGOALS(Blast_tac)); |
3656 | 445 |
qed"is_asig_of_rename"; |
3521 | 446 |
|
447 |
||
3656 | 448 |
Addsimps [is_asig_of_par,is_asig_of_restrict,is_asig_of_rename, |
449 |
is_trans_of_par,is_trans_of_restrict,is_trans_of_rename]; |
|
3521 | 450 |
|
451 |
||
5068 | 452 |
Goalw [compatible_def] |
6161 | 453 |
"[|compatible A B; compatible A C |]==> compatible A (B||C)"; |
4098 | 454 |
by (asm_full_simp_tac (simpset() addsimps [internals_of_par, |
3656 | 455 |
outputs_of_par,actions_of_par]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
456 |
by Auto_tac; |
3656 | 457 |
qed"compatible_par"; |
3521 | 458 |
|
5976 | 459 |
(* better derive by previous one and compat_commute *) |
5068 | 460 |
Goalw [compatible_def] |
6161 | 461 |
"[|compatible A C; compatible B C |]==> compatible (A||B) C"; |
4098 | 462 |
by (asm_full_simp_tac (simpset() addsimps [internals_of_par, |
3656 | 463 |
outputs_of_par,actions_of_par]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
464 |
by Auto_tac; |
3656 | 465 |
qed"compatible_par2"; |
3521 | 466 |
|
5068 | 467 |
Goalw [compatible_def] |
6161 | 468 |
"[| compatible A B; (ext B - S) Int ext A = {}|] \ |
3656 | 469 |
\ ==> compatible A (restrict B S)"; |
4098 | 470 |
by (asm_full_simp_tac (simpset() addsimps [ioa_triple_proj,asig_triple_proj, |
3656 | 471 |
externals_def,restrict_def,restrict_asig_def,actions_def]) 1); |
6161 | 472 |
by Auto_tac; |
3656 | 473 |
qed"compatible_restrict"; |
4814 | 474 |
|
475 |
||
476 |
Addsimps [split_paired_Ex]; |