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