| author | haftmann |
| Tue, 05 Jul 2005 13:57:23 +0200 | |
| changeset 16688 | e3de7ea24c23 |
| parent 14981 | e73f8140af78 |
| child 17233 | 41eee2e7b465 |
| permissions | -rw-r--r-- |
| 3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/Compositionality.ML |
| 3275 | 2 |
ID: $Id$ |
| 12218 | 3 |
Author: Olaf Müller |
| 3071 | 4 |
|
5 |
Compositionality of I/O automata |
|
6 |
*) |
|
7 |
||
8 |
||
| 3275 | 9 |
|
| 6161 | 10 |
Goal "[|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
11 |
by Auto_tac; |
| 3275 | 12 |
qed"compatibility_consequence3"; |
13 |
||
14 |
||
| 5068 | 15 |
Goal |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
16 |
"!! A B. [| compatible A B; Forall (%a. a: ext A | a: ext B) tr |] ==> \ |
| 10835 | 17 |
\ Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"; |
| 3457 | 18 |
by (rtac ForallPFilterQR 1); |
| 10835 | 19 |
(* i.e.: [| (! x. P x --> (Q x = R x)) ; Forall P tr |] ==> Filter Q$tr = Filter R$tr *) |
| 3457 | 20 |
by (assume_tac 2); |
21 |
by (rtac compatibility_consequence3 1); |
|
| 4098 | 22 |
by (REPEAT (asm_full_simp_tac (simpset() |
| 3275 | 23 |
addsimps [ext_is_act,ext1_ext2_is_not_act1]) 1)); |
24 |
qed"Filter_actAisFilter_extA"; |
|
25 |
||
26 |
||
| 5976 | 27 |
(* the next two theorems are only necessary, as there is no theorem ext (A||B) = ext (B||A) *) |
| 3275 | 28 |
|
| 6161 | 29 |
Goal "[|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
30 |
by Auto_tac; |
| 3275 | 31 |
qed"compatibility_consequence4"; |
32 |
||
| 6161 | 33 |
Goal "[| compatible A B; Forall (%a. a: ext B | a: ext A) tr |] ==> \ |
| 10835 | 34 |
\ Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"; |
| 3457 | 35 |
by (rtac ForallPFilterQR 1); |
36 |
by (assume_tac 2); |
|
37 |
by (rtac compatibility_consequence4 1); |
|
| 4098 | 38 |
by (REPEAT (asm_full_simp_tac (simpset() |
| 3275 | 39 |
addsimps [ext_is_act,ext1_ext2_is_not_act1]) 1)); |
40 |
qed"Filter_actAisFilter_extA2"; |
|
41 |
||
42 |
||
43 |
(* -------------------------------------------------------------------------- *) |
|
44 |
section " Main Compositionality Theorem "; |
|
45 |
(* -------------------------------------------------------------------------- *) |
|
46 |
||
47 |
||
48 |
||
| 6161 | 49 |
Goal "[| is_trans_of A1; is_trans_of A2; is_trans_of B1; is_trans_of B2;\ |
| 3521 | 50 |
\ is_asig_of A1; is_asig_of A2; \ |
51 |
\ is_asig_of B1; is_asig_of B2; \ |
|
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
52 |
\ compatible A1 B1; compatible A2 B2; \ |
| 3071 | 53 |
\ A1 =<| A2; \ |
54 |
\ B1 =<| B2 |] \ |
|
55 |
\ ==> (A1 || B1) =<| (A2 || B2)"; |
|
| 3275 | 56 |
|
| 4098 | 57 |
by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def]) 1); |
| 4473 | 58 |
by (forw_inst_tac [("A1","A1")] (compat_commute RS iffD1) 1);
|
59 |
by (forw_inst_tac [("A1","A2")] (compat_commute RS iffD1) 1);
|
|
| 4098 | 60 |
by (asm_full_simp_tac (simpset() addsimps [ioa_implements_def, |
| 3275 | 61 |
inputs_of_par,outputs_of_par,externals_of_par]) 1); |
62 |
by (safe_tac set_cs); |
|
| 4098 | 63 |
by (asm_full_simp_tac (simpset() addsimps [compositionality_tr]) 1); |
| 3275 | 64 |
by (subgoal_tac "ext A1 = ext A2 & ext B1 = ext B2" 1); |
| 4098 | 65 |
by (asm_full_simp_tac (simpset() addsimps [externals_def]) 2); |
| 3275 | 66 |
by (REPEAT (etac conjE 1)); |
67 |
(* rewrite with proven subgoal *) |
|
| 4098 | 68 |
by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1); |
| 3275 | 69 |
by (safe_tac set_cs); |
70 |
||
71 |
(* 2 goals, the 3rd has been solved automatically *) |
|
72 |
(* 1: Filter A2 x : traces A2 *) |
|
73 |
by (dres_inst_tac [("A","traces A1")] subsetD 1);
|
|
| 3457 | 74 |
by (assume_tac 1); |
| 4098 | 75 |
by (asm_full_simp_tac (simpset() addsimps [Filter_actAisFilter_extA]) 1); |
| 3275 | 76 |
(* 2: Filter B2 x : traces B2 *) |
77 |
by (dres_inst_tac [("A","traces B1")] subsetD 1);
|
|
| 3457 | 78 |
by (assume_tac 1); |
| 4098 | 79 |
by (asm_full_simp_tac (simpset() addsimps [Filter_actAisFilter_extA2]) 1); |
| 3275 | 80 |
qed"compositionality"; |
| 3071 | 81 |
|
82 |
||
83 |
||
| 3275 | 84 |