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