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