| 3071 |      1 | (*  Title:      HOLCF/IOA/meta_theory/CompoScheds.thy
 | 
| 3275 |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 17233 |      4 | *)
 | 
| 3071 |      5 | 
 | 
| 17233 |      6 | header {* Compositionality on Schedule level *}
 | 
| 3071 |      7 | 
 | 
| 17233 |      8 | theory CompoScheds
 | 
|  |      9 | imports CompoExecs
 | 
|  |     10 | begin
 | 
| 3071 |     11 | 
 | 
|  |     12 | consts
 | 
|  |     13 | 
 | 
|  |     14 |  mkex     ::"('a,'s)ioa => ('a,'t)ioa => 'a Seq =>
 | 
| 17233 |     15 |               ('a,'s)execution => ('a,'t)execution =>('a,'s*'t)execution"
 | 
|  |     16 |  mkex2    ::"('a,'s)ioa => ('a,'t)ioa => 'a Seq ->
 | 
|  |     17 |               ('a,'s)pairs -> ('a,'t)pairs ->
 | 
| 3071 |     18 |               ('s => 't => ('a,'s*'t)pairs)"
 | 
| 3521 |     19 |  par_scheds ::"['a schedule_module,'a schedule_module] => 'a schedule_module"
 | 
|  |     20 | 
 | 
| 3071 |     21 | 
 | 
|  |     22 | defs
 | 
|  |     23 | 
 | 
| 17233 |     24 | mkex_def:
 | 
|  |     25 |   "mkex A B sch exA exB ==
 | 
| 3071 |     26 |        ((fst exA,fst exB),
 | 
| 10835 |     27 |         (mkex2 A B$sch$(snd exA)$(snd exB)) (fst exA) (fst exB))"
 | 
| 3071 |     28 | 
 | 
| 17233 |     29 | mkex2_def:
 | 
|  |     30 |   "mkex2 A B == (fix$(LAM h sch exA exB. (%s t. case sch of
 | 
| 3071 |     31 |        nil => nil
 | 
| 17233 |     32 |     | x##xs =>
 | 
|  |     33 |       (case x of
 | 
| 12028 |     34 |         UU => UU
 | 
| 17233 |     35 |       | Def y =>
 | 
|  |     36 |          (if y:act A then
 | 
|  |     37 |              (if y:act B then
 | 
| 10835 |     38 |                 (case HD$exA of
 | 
| 12028 |     39 |                    UU => UU
 | 
| 10835 |     40 |                  | Def a => (case HD$exB of
 | 
| 12028 |     41 |                               UU => UU
 | 
| 17233 |     42 |                             | Def b =>
 | 
| 3071 |     43 |                    (y,(snd a,snd b))>>
 | 
| 10835 |     44 |                      (h$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))
 | 
| 3071 |     45 |               else
 | 
| 10835 |     46 |                 (case HD$exA of
 | 
| 12028 |     47 |                    UU => UU
 | 
| 17233 |     48 |                  | Def a =>
 | 
| 10835 |     49 |                    (y,(snd a,t))>>(h$xs$(TL$exA)$exB) (snd a) t)
 | 
| 3071 |     50 |               )
 | 
| 17233 |     51 |           else
 | 
|  |     52 |              (if y:act B then
 | 
| 10835 |     53 |                 (case HD$exB of
 | 
| 12028 |     54 |                    UU => UU
 | 
| 17233 |     55 |                  | Def b =>
 | 
| 10835 |     56 |                    (y,(s,snd b))>>(h$xs$exA$(TL$exB)) s (snd b))
 | 
| 3071 |     57 |              else
 | 
|  |     58 |                UU
 | 
|  |     59 |              )
 | 
|  |     60 |          )
 | 
|  |     61 |        ))))"
 | 
|  |     62 | 
 | 
| 17233 |     63 | par_scheds_def:
 | 
|  |     64 |   "par_scheds SchedsA SchedsB ==
 | 
|  |     65 |        let schA = fst SchedsA; sigA = snd SchedsA;
 | 
|  |     66 |            schB = fst SchedsB; sigB = snd SchedsB
 | 
| 3521 |     67 |        in
 | 
| 10835 |     68 |        (    {sch. Filter (%a. a:actions sigA)$sch : schA}
 | 
|  |     69 |         Int {sch. Filter (%a. a:actions sigB)$sch : schB}
 | 
| 3521 |     70 |         Int {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch},
 | 
|  |     71 |         asig_comp sigA sigB)"
 | 
|  |     72 | 
 | 
| 19741 |     73 | 
 | 
|  |     74 | declare surjective_pairing [symmetric, simp]
 | 
|  |     75 | 
 | 
|  |     76 | 
 | 
|  |     77 | subsection "mkex rewrite rules"
 | 
|  |     78 | 
 | 
|  |     79 | 
 | 
|  |     80 | lemma mkex2_unfold:
 | 
|  |     81 | "mkex2 A B = (LAM sch exA exB. (%s t. case sch of  
 | 
|  |     82 |       nil => nil 
 | 
|  |     83 |    | x##xs => 
 | 
|  |     84 |      (case x of  
 | 
|  |     85 |        UU => UU   
 | 
|  |     86 |      | Def y =>  
 | 
|  |     87 |         (if y:act A then 
 | 
|  |     88 |             (if y:act B then 
 | 
|  |     89 |                (case HD$exA of 
 | 
|  |     90 |                   UU => UU 
 | 
|  |     91 |                 | Def a => (case HD$exB of 
 | 
|  |     92 |                              UU => UU 
 | 
|  |     93 |                            | Def b => 
 | 
|  |     94 |                   (y,(snd a,snd b))>>  
 | 
|  |     95 |                     (mkex2 A B$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))  
 | 
|  |     96 |              else   
 | 
|  |     97 |                (case HD$exA of   
 | 
|  |     98 |                   UU => UU  
 | 
|  |     99 |                 | Def a => 
 | 
|  |    100 |                   (y,(snd a,t))>>(mkex2 A B$xs$(TL$exA)$exB) (snd a) t)  
 | 
|  |    101 |              )  
 | 
|  |    102 |          else   
 | 
|  |    103 |             (if y:act B then 
 | 
|  |    104 |                (case HD$exB of  
 | 
|  |    105 |                   UU => UU  
 | 
|  |    106 |                 | Def b =>  
 | 
|  |    107 |                   (y,(s,snd b))>>(mkex2 A B$xs$exA$(TL$exB)) s (snd b))  
 | 
|  |    108 |             else  
 | 
|  |    109 |               UU  
 | 
|  |    110 |             )  
 | 
|  |    111 |         )  
 | 
|  |    112 |       )))"
 | 
|  |    113 | apply (rule trans)
 | 
|  |    114 | apply (rule fix_eq2)
 | 
|  |    115 | apply (rule mkex2_def)
 | 
|  |    116 | apply (rule beta_cfun)
 | 
|  |    117 | apply simp
 | 
|  |    118 | done
 | 
|  |    119 | 
 | 
|  |    120 | lemma mkex2_UU: "(mkex2 A B$UU$exA$exB) s t = UU"
 | 
|  |    121 | apply (subst mkex2_unfold)
 | 
|  |    122 | apply simp
 | 
|  |    123 | done
 | 
|  |    124 | 
 | 
|  |    125 | lemma mkex2_nil: "(mkex2 A B$nil$exA$exB) s t= nil"
 | 
|  |    126 | apply (subst mkex2_unfold)
 | 
|  |    127 | apply simp
 | 
|  |    128 | done
 | 
|  |    129 | 
 | 
|  |    130 | lemma mkex2_cons_1: "[| x:act A; x~:act B; HD$exA=Def a|]  
 | 
|  |    131 |     ==> (mkex2 A B$(x>>sch)$exA$exB) s t =    
 | 
|  |    132 |         (x,snd a,t) >> (mkex2 A B$sch$(TL$exA)$exB) (snd a) t"
 | 
|  |    133 | apply (rule trans)
 | 
|  |    134 | apply (subst mkex2_unfold)
 | 
|  |    135 | apply (simp add: Consq_def If_and_if)
 | 
|  |    136 | apply (simp add: Consq_def)
 | 
|  |    137 | done
 | 
|  |    138 | 
 | 
|  |    139 | lemma mkex2_cons_2: "[| x~:act A; x:act B; HD$exB=Def b|]  
 | 
|  |    140 |     ==> (mkex2 A B$(x>>sch)$exA$exB) s t =  
 | 
|  |    141 |         (x,s,snd b) >> (mkex2 A B$sch$exA$(TL$exB)) s (snd b)"
 | 
|  |    142 | apply (rule trans)
 | 
|  |    143 | apply (subst mkex2_unfold)
 | 
|  |    144 | apply (simp add: Consq_def If_and_if)
 | 
|  |    145 | apply (simp add: Consq_def)
 | 
|  |    146 | done
 | 
|  |    147 | 
 | 
|  |    148 | lemma mkex2_cons_3: "[| x:act A; x:act B; HD$exA=Def a;HD$exB=Def b|]  
 | 
|  |    149 |     ==> (mkex2 A B$(x>>sch)$exA$exB) s t =   
 | 
|  |    150 |          (x,snd a,snd b) >>  
 | 
|  |    151 |             (mkex2 A B$sch$(TL$exA)$(TL$exB)) (snd a) (snd b)"
 | 
|  |    152 | apply (rule trans)
 | 
|  |    153 | apply (subst mkex2_unfold)
 | 
|  |    154 | apply (simp add: Consq_def If_and_if)
 | 
|  |    155 | apply (simp add: Consq_def)
 | 
|  |    156 | done
 | 
|  |    157 | 
 | 
|  |    158 | declare mkex2_UU [simp] mkex2_nil [simp] mkex2_cons_1 [simp]
 | 
|  |    159 |   mkex2_cons_2 [simp] mkex2_cons_3 [simp]
 | 
|  |    160 | 
 | 
|  |    161 | 
 | 
|  |    162 | subsection {* mkex *}
 | 
|  |    163 | 
 | 
|  |    164 | lemma mkex_UU: "mkex A B UU  (s,exA) (t,exB) = ((s,t),UU)"
 | 
|  |    165 | apply (simp add: mkex_def)
 | 
|  |    166 | done
 | 
|  |    167 | 
 | 
|  |    168 | lemma mkex_nil: "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)"
 | 
|  |    169 | apply (simp add: mkex_def)
 | 
|  |    170 | done
 | 
|  |    171 | 
 | 
|  |    172 | lemma mkex_cons_1: "[| x:act A; x~:act B |]  
 | 
|  |    173 |     ==> mkex A B (x>>sch) (s,a>>exA) (t,exB)  =   
 | 
|  |    174 |         ((s,t), (x,snd a,t) >> snd (mkex A B sch (snd a,exA) (t,exB)))"
 | 
|  |    175 | apply (simp (no_asm) add: mkex_def)
 | 
|  |    176 | apply (cut_tac exA = "a>>exA" in mkex2_cons_1)
 | 
|  |    177 | apply auto
 | 
|  |    178 | done
 | 
|  |    179 | 
 | 
|  |    180 | lemma mkex_cons_2: "[| x~:act A; x:act B |]  
 | 
|  |    181 |     ==> mkex A B (x>>sch) (s,exA) (t,b>>exB) =   
 | 
|  |    182 |         ((s,t), (x,s,snd b) >> snd (mkex A B sch (s,exA) (snd b,exB)))"
 | 
|  |    183 | apply (simp (no_asm) add: mkex_def)
 | 
|  |    184 | apply (cut_tac exB = "b>>exB" in mkex2_cons_2)
 | 
|  |    185 | apply auto
 | 
|  |    186 | done
 | 
|  |    187 | 
 | 
|  |    188 | lemma mkex_cons_3: "[| x:act A; x:act B |]   
 | 
|  |    189 |     ==>  mkex A B (x>>sch) (s,a>>exA) (t,b>>exB) =  
 | 
|  |    190 |          ((s,t), (x,snd a,snd b) >> snd (mkex A B sch (snd a,exA) (snd b,exB)))"
 | 
|  |    191 | apply (simp (no_asm) add: mkex_def)
 | 
|  |    192 | apply (cut_tac exB = "b>>exB" and exA = "a>>exA" in mkex2_cons_3)
 | 
|  |    193 | apply auto
 | 
|  |    194 | done
 | 
|  |    195 | 
 | 
|  |    196 | declare mkex2_UU [simp del] mkex2_nil [simp del]
 | 
|  |    197 |   mkex2_cons_1 [simp del] mkex2_cons_2 [simp del] mkex2_cons_3 [simp del]
 | 
|  |    198 | 
 | 
|  |    199 | lemmas composch_simps = mkex_UU mkex_nil mkex_cons_1 mkex_cons_2 mkex_cons_3
 | 
|  |    200 | 
 | 
|  |    201 | declare composch_simps [simp]
 | 
|  |    202 | 
 | 
|  |    203 | 
 | 
|  |    204 | subsection {* COMPOSITIONALITY on SCHEDULE Level *}
 | 
|  |    205 | 
 | 
|  |    206 | subsubsection "Lemmas for ==>"
 | 
|  |    207 | 
 | 
|  |    208 | (* --------------------------------------------------------------------- *)
 | 
|  |    209 | (*    Lemma_2_1 :  tfilter(ex) and filter_act are commutative            *)
 | 
|  |    210 | (* --------------------------------------------------------------------- *)
 | 
|  |    211 | 
 | 
|  |    212 | lemma lemma_2_1a: 
 | 
|  |    213 |    "filter_act$(Filter_ex2 (asig_of A)$xs)= 
 | 
|  |    214 |     Filter (%a. a:act A)$(filter_act$xs)"
 | 
|  |    215 | 
 | 
|  |    216 | apply (unfold filter_act_def Filter_ex2_def)
 | 
|  |    217 | apply (simp (no_asm) add: MapFilter o_def)
 | 
|  |    218 | done
 | 
|  |    219 | 
 | 
|  |    220 | 
 | 
|  |    221 | (* --------------------------------------------------------------------- *)
 | 
|  |    222 | (*    Lemma_2_2 : State-projections do not affect filter_act             *)
 | 
|  |    223 | (* --------------------------------------------------------------------- *)
 | 
|  |    224 | 
 | 
|  |    225 | lemma lemma_2_1b: 
 | 
|  |    226 |    "filter_act$(ProjA2$xs) =filter_act$xs & 
 | 
|  |    227 |     filter_act$(ProjB2$xs) =filter_act$xs"
 | 
|  |    228 | apply (tactic {* pair_induct_tac "xs" [] 1 *})
 | 
|  |    229 | done
 | 
|  |    230 | 
 | 
|  |    231 | 
 | 
|  |    232 | (* --------------------------------------------------------------------- *)
 | 
|  |    233 | (*             Schedules of A||B have only  A- or B-actions              *)
 | 
|  |    234 | (* --------------------------------------------------------------------- *)
 | 
|  |    235 | 
 | 
|  |    236 | (* very similar to lemma_1_1c, but it is not checking if every action element of
 | 
|  |    237 |    an ex is in A or B, but after projecting it onto the action schedule. Of course, this
 | 
|  |    238 |    is the same proposition, but we cannot change this one, when then rather lemma_1_1c  *)
 | 
|  |    239 | 
 | 
|  |    240 | lemma sch_actions_in_AorB: "!s. is_exec_frag (A||B) (s,xs)  
 | 
|  |    241 |    --> Forall (%x. x:act (A||B)) (filter_act$xs)"
 | 
|  |    242 | 
 | 
|  |    243 | apply (tactic {* pair_induct_tac "xs" [thm "is_exec_frag_def", thm "Forall_def",
 | 
|  |    244 |   thm "sforall_def"] 1 *})
 | 
|  |    245 | (* main case *)
 | 
|  |    246 | apply (tactic "safe_tac set_cs")
 | 
|  |    247 | apply (auto simp add: trans_of_defs2 actions_asig_comp asig_of_par)
 | 
|  |    248 | done
 | 
|  |    249 | 
 | 
|  |    250 | 
 | 
|  |    251 | subsubsection "Lemmas for <=="
 | 
|  |    252 | 
 | 
|  |    253 | (*---------------------------------------------------------------------------
 | 
|  |    254 |     Filtering actions out of mkex(sch,exA,exB) yields the oracle sch
 | 
|  |    255 |                              structural induction
 | 
|  |    256 |   --------------------------------------------------------------------------- *)
 | 
|  |    257 | 
 | 
|  |    258 | lemma Mapfst_mkex_is_sch: "! exA exB s t.  
 | 
|  |    259 |   Forall (%x. x:act (A||B)) sch  &  
 | 
|  |    260 |   Filter (%a. a:act A)$sch << filter_act$exA & 
 | 
|  |    261 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    262 |   --> filter_act$(snd (mkex A B sch (s,exA) (t,exB))) = sch"
 | 
|  |    263 | 
 | 
|  |    264 | apply (tactic {* Seq_induct_tac "sch" [thm "Filter_def", thm "Forall_def",
 | 
|  |    265 |   thm "sforall_def", thm "mkex_def"] 1 *})
 | 
|  |    266 | 
 | 
|  |    267 | (* main case *)
 | 
|  |    268 | (* splitting into 4 cases according to a:A, a:B *)
 | 
|  |    269 | apply (tactic "safe_tac set_cs")
 | 
|  |    270 | 
 | 
|  |    271 | (* Case y:A, y:B *)
 | 
|  |    272 | apply (tactic {* Seq_case_simp_tac "exA" 1 *})
 | 
|  |    273 | (* Case exA=UU, Case exA=nil*)
 | 
|  |    274 | (* These UU and nil cases are the only places where the assumption filter A sch<<f_act exA
 | 
|  |    275 |    is used! --> to generate a contradiction using  ~a>>ss<< UU(nil), using theorems
 | 
|  |    276 |    Cons_not_less_UU and Cons_not_less_nil  *)
 | 
|  |    277 | apply (tactic {* Seq_case_simp_tac "exB" 1 *})
 | 
|  |    278 | (* Case exA=a>>x, exB=b>>y *)
 | 
|  |    279 | (* here it is important that Seq_case_simp_tac uses no !full!_simp_tac for the cons case,
 | 
|  |    280 |    as otherwise mkex_cons_3 would  not be rewritten without use of rotate_tac: then tactic
 | 
|  |    281 |    would not be generally applicable *)
 | 
|  |    282 | apply simp
 | 
|  |    283 | 
 | 
|  |    284 | (* Case y:A, y~:B *)
 | 
|  |    285 | apply (tactic {* Seq_case_simp_tac "exA" 1 *})
 | 
|  |    286 | apply simp
 | 
|  |    287 | 
 | 
|  |    288 | (* Case y~:A, y:B *)
 | 
|  |    289 | apply (tactic {* Seq_case_simp_tac "exB" 1 *})
 | 
|  |    290 | apply simp
 | 
|  |    291 | 
 | 
|  |    292 | (* Case y~:A, y~:B *)
 | 
|  |    293 | apply (simp add: asig_of_par actions_asig_comp)
 | 
|  |    294 | done
 | 
|  |    295 | 
 | 
|  |    296 | 
 | 
|  |    297 | (* generalizing the proof above to a tactic *)
 | 
|  |    298 | 
 | 
|  |    299 | ML {*
 | 
|  |    300 | 
 | 
|  |    301 | local
 | 
|  |    302 |   val defs = [thm "Filter_def", thm "Forall_def", thm "sforall_def", thm "mkex_def",
 | 
|  |    303 |     thm "stutter_def"]
 | 
|  |    304 |   val asigs = [thm "asig_of_par", thm "actions_asig_comp"]
 | 
|  |    305 | in
 | 
|  |    306 | 
 | 
|  |    307 | fun mkex_induct_tac sch exA exB =
 | 
|  |    308 |     EVERY1[Seq_induct_tac sch defs,
 | 
|  |    309 |            SIMPSET' asm_full_simp_tac,
 | 
|  |    310 |            SELECT_GOAL (safe_tac set_cs),
 | 
|  |    311 |            Seq_case_simp_tac exA,
 | 
|  |    312 |            Seq_case_simp_tac exB,
 | 
|  |    313 |            SIMPSET' asm_full_simp_tac,
 | 
|  |    314 |            Seq_case_simp_tac exA,
 | 
|  |    315 |            SIMPSET' asm_full_simp_tac,
 | 
|  |    316 |            Seq_case_simp_tac exB,
 | 
|  |    317 |            SIMPSET' asm_full_simp_tac,
 | 
|  |    318 |            SIMPSET' (fn ss => asm_full_simp_tac (ss addsimps asigs))
 | 
|  |    319 |           ]
 | 
| 17233 |    320 | 
 | 
| 3521 |    321 | end
 | 
| 19741 |    322 | *}
 | 
|  |    323 | 
 | 
|  |    324 | 
 | 
|  |    325 | (*---------------------------------------------------------------------------
 | 
|  |    326 |                Projection of mkex(sch,exA,exB) onto A stutters on A
 | 
|  |    327 |                              structural induction
 | 
|  |    328 |   --------------------------------------------------------------------------- *)
 | 
|  |    329 | 
 | 
|  |    330 | lemma stutterA_mkex: "! exA exB s t.  
 | 
|  |    331 |   Forall (%x. x:act (A||B)) sch &  
 | 
|  |    332 |   Filter (%a. a:act A)$sch << filter_act$exA & 
 | 
|  |    333 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    334 |   --> stutter (asig_of A) (s,ProjA2$(snd (mkex A B sch (s,exA) (t,exB))))"
 | 
|  |    335 | 
 | 
|  |    336 | apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
 | 
|  |    337 | done
 | 
|  |    338 | 
 | 
|  |    339 | 
 | 
|  |    340 | lemma stutter_mkex_on_A: "[|   
 | 
|  |    341 |   Forall (%x. x:act (A||B)) sch ;  
 | 
|  |    342 |   Filter (%a. a:act A)$sch << filter_act$(snd exA) ; 
 | 
|  |    343 |   Filter (%a. a:act B)$sch << filter_act$(snd exB) |]  
 | 
|  |    344 |   ==> stutter (asig_of A) (ProjA (mkex A B sch exA exB))"
 | 
|  |    345 | 
 | 
|  |    346 | apply (cut_tac stutterA_mkex)
 | 
|  |    347 | apply (simp add: stutter_def ProjA_def mkex_def)
 | 
|  |    348 | apply (erule allE)+
 | 
|  |    349 | apply (drule mp)
 | 
|  |    350 | prefer 2 apply (assumption)
 | 
|  |    351 | apply simp
 | 
|  |    352 | done
 | 
|  |    353 | 
 | 
|  |    354 | 
 | 
|  |    355 | (*---------------------------------------------------------------------------
 | 
|  |    356 |                Projection of mkex(sch,exA,exB) onto B stutters on B
 | 
|  |    357 |                              structural induction
 | 
|  |    358 |   --------------------------------------------------------------------------- *)
 | 
|  |    359 | 
 | 
|  |    360 | lemma stutterB_mkex: "! exA exB s t.  
 | 
|  |    361 |   Forall (%x. x:act (A||B)) sch &  
 | 
|  |    362 |   Filter (%a. a:act A)$sch << filter_act$exA & 
 | 
|  |    363 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    364 |   --> stutter (asig_of B) (t,ProjB2$(snd (mkex A B sch (s,exA) (t,exB))))"
 | 
|  |    365 | apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
 | 
|  |    366 | done
 | 
|  |    367 | 
 | 
|  |    368 | 
 | 
|  |    369 | lemma stutter_mkex_on_B: "[|   
 | 
|  |    370 |   Forall (%x. x:act (A||B)) sch ;  
 | 
|  |    371 |   Filter (%a. a:act A)$sch << filter_act$(snd exA) ; 
 | 
|  |    372 |   Filter (%a. a:act B)$sch << filter_act$(snd exB) |]  
 | 
|  |    373 |   ==> stutter (asig_of B) (ProjB (mkex A B sch exA exB))"
 | 
|  |    374 | apply (cut_tac stutterB_mkex)
 | 
|  |    375 | apply (simp add: stutter_def ProjB_def mkex_def)
 | 
|  |    376 | apply (erule allE)+
 | 
|  |    377 | apply (drule mp)
 | 
|  |    378 | prefer 2 apply (assumption)
 | 
|  |    379 | apply simp
 | 
|  |    380 | done
 | 
|  |    381 | 
 | 
|  |    382 | 
 | 
|  |    383 | (*---------------------------------------------------------------------------
 | 
|  |    384 |      Filter of mkex(sch,exA,exB) to A after projection onto A is exA
 | 
|  |    385 |         --  using zip$(proj1$exA)$(proj2$exA) instead of exA    --
 | 
|  |    386 |         --           because of admissibility problems          --
 | 
|  |    387 |                              structural induction
 | 
|  |    388 |   --------------------------------------------------------------------------- *)
 | 
|  |    389 | 
 | 
|  |    390 | lemma filter_mkex_is_exA_tmp: "! exA exB s t.  
 | 
|  |    391 |   Forall (%x. x:act (A||B)) sch &  
 | 
|  |    392 |   Filter (%a. a:act A)$sch << filter_act$exA  & 
 | 
|  |    393 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    394 |   --> Filter_ex2 (asig_of A)$(ProjA2$(snd (mkex A B sch (s,exA) (t,exB)))) =    
 | 
|  |    395 |       Zip$(Filter (%a. a:act A)$sch)$(Map snd$exA)"
 | 
|  |    396 | apply (tactic {* mkex_induct_tac "sch" "exB" "exA" *})
 | 
|  |    397 | done
 | 
|  |    398 | 
 | 
|  |    399 | (*---------------------------------------------------------------------------
 | 
|  |    400 |                       zip$(proj1$y)$(proj2$y) = y   (using the lift operations)
 | 
|  |    401 |                     lemma for admissibility problems
 | 
|  |    402 |   --------------------------------------------------------------------------- *)
 | 
|  |    403 | 
 | 
|  |    404 | lemma Zip_Map_fst_snd: "Zip$(Map fst$y)$(Map snd$y) = y"
 | 
|  |    405 | apply (tactic {* Seq_induct_tac "y" [] 1 *})
 | 
|  |    406 | done
 | 
|  |    407 | 
 | 
|  |    408 | 
 | 
|  |    409 | (*---------------------------------------------------------------------------
 | 
|  |    410 |       filter A$sch = proj1$ex   -->  zip$(filter A$sch)$(proj2$ex) = ex
 | 
|  |    411 |          lemma for eliminating non admissible equations in assumptions
 | 
|  |    412 |   --------------------------------------------------------------------------- *)
 | 
|  |    413 | 
 | 
|  |    414 | lemma trick_against_eq_in_ass: "!! sch ex.  
 | 
|  |    415 |   Filter (%a. a:act AB)$sch = filter_act$ex   
 | 
|  |    416 |   ==> ex = Zip$(Filter (%a. a:act AB)$sch)$(Map snd$ex)"
 | 
|  |    417 | apply (simp add: filter_act_def)
 | 
|  |    418 | apply (rule Zip_Map_fst_snd [symmetric])
 | 
|  |    419 | done
 | 
|  |    420 | 
 | 
|  |    421 | (*---------------------------------------------------------------------------
 | 
|  |    422 |      Filter of mkex(sch,exA,exB) to A after projection onto A is exA
 | 
|  |    423 |                        using the above trick
 | 
|  |    424 |   --------------------------------------------------------------------------- *)
 | 
|  |    425 | 
 | 
|  |    426 | 
 | 
|  |    427 | lemma filter_mkex_is_exA: "!!sch exA exB. 
 | 
|  |    428 |   [| Forall (%a. a:act (A||B)) sch ;  
 | 
|  |    429 |   Filter (%a. a:act A)$sch = filter_act$(snd exA)  ; 
 | 
|  |    430 |   Filter (%a. a:act B)$sch = filter_act$(snd exB) |] 
 | 
|  |    431 |   ==> Filter_ex (asig_of A) (ProjA (mkex A B sch exA exB)) = exA"
 | 
|  |    432 | apply (simp add: ProjA_def Filter_ex_def)
 | 
|  |    433 | apply (tactic {* pair_tac "exA" 1 *})
 | 
|  |    434 | apply (tactic {* pair_tac "exB" 1 *})
 | 
|  |    435 | apply (rule conjI)
 | 
|  |    436 | apply (simp (no_asm) add: mkex_def)
 | 
|  |    437 | apply (simplesubst trick_against_eq_in_ass)
 | 
|  |    438 | back
 | 
|  |    439 | apply assumption
 | 
|  |    440 | apply (simp add: filter_mkex_is_exA_tmp)
 | 
|  |    441 | done
 | 
|  |    442 | 
 | 
|  |    443 | 
 | 
|  |    444 | (*---------------------------------------------------------------------------
 | 
|  |    445 |      Filter of mkex(sch,exA,exB) to B after projection onto B is exB
 | 
|  |    446 |         --  using zip$(proj1$exB)$(proj2$exB) instead of exB    --
 | 
|  |    447 |         --           because of admissibility problems          --
 | 
|  |    448 |                              structural induction
 | 
|  |    449 |   --------------------------------------------------------------------------- *)
 | 
|  |    450 | 
 | 
|  |    451 | lemma filter_mkex_is_exB_tmp: "! exA exB s t.  
 | 
|  |    452 |   Forall (%x. x:act (A||B)) sch &  
 | 
|  |    453 |   Filter (%a. a:act A)$sch << filter_act$exA  & 
 | 
|  |    454 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    455 |   --> Filter_ex2 (asig_of B)$(ProjB2$(snd (mkex A B sch (s,exA) (t,exB)))) =    
 | 
|  |    456 |       Zip$(Filter (%a. a:act B)$sch)$(Map snd$exB)"
 | 
|  |    457 | 
 | 
|  |    458 | (* notice necessary change of arguments exA and exB *)
 | 
|  |    459 | apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
 | 
|  |    460 | done
 | 
|  |    461 | 
 | 
|  |    462 | 
 | 
|  |    463 | (*---------------------------------------------------------------------------
 | 
|  |    464 |      Filter of mkex(sch,exA,exB) to A after projection onto B is exB
 | 
|  |    465 |                        using the above trick
 | 
|  |    466 |   --------------------------------------------------------------------------- *)
 | 
|  |    467 | 
 | 
|  |    468 | 
 | 
|  |    469 | lemma filter_mkex_is_exB: "!!sch exA exB. 
 | 
|  |    470 |   [| Forall (%a. a:act (A||B)) sch ;  
 | 
|  |    471 |   Filter (%a. a:act A)$sch = filter_act$(snd exA)  ; 
 | 
|  |    472 |   Filter (%a. a:act B)$sch = filter_act$(snd exB) |] 
 | 
|  |    473 |   ==> Filter_ex (asig_of B) (ProjB (mkex A B sch exA exB)) = exB"
 | 
|  |    474 | apply (simp add: ProjB_def Filter_ex_def)
 | 
|  |    475 | apply (tactic {* pair_tac "exA" 1 *})
 | 
|  |    476 | apply (tactic {* pair_tac "exB" 1 *})
 | 
|  |    477 | apply (rule conjI)
 | 
|  |    478 | apply (simp (no_asm) add: mkex_def)
 | 
|  |    479 | apply (simplesubst trick_against_eq_in_ass)
 | 
|  |    480 | back
 | 
|  |    481 | apply assumption
 | 
|  |    482 | apply (simp add: filter_mkex_is_exB_tmp)
 | 
|  |    483 | done
 | 
|  |    484 | 
 | 
|  |    485 | (* --------------------------------------------------------------------- *)
 | 
|  |    486 | (*                    mkex has only  A- or B-actions                    *)
 | 
|  |    487 | (* --------------------------------------------------------------------- *)
 | 
|  |    488 | 
 | 
|  |    489 | 
 | 
|  |    490 | lemma mkex_actions_in_AorB: "!s t exA exB.  
 | 
|  |    491 |   Forall (%x. x : act (A || B)) sch & 
 | 
|  |    492 |   Filter (%a. a:act A)$sch << filter_act$exA  & 
 | 
|  |    493 |   Filter (%a. a:act B)$sch << filter_act$exB  
 | 
|  |    494 |    --> Forall (%x. fst x : act (A ||B))    
 | 
|  |    495 |          (snd (mkex A B sch (s,exA) (t,exB)))"
 | 
|  |    496 | apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
 | 
|  |    497 | done
 | 
|  |    498 | 
 | 
|  |    499 | 
 | 
|  |    500 | (* ------------------------------------------------------------------ *)
 | 
|  |    501 | (*           COMPOSITIONALITY   on    SCHEDULE      Level             *)
 | 
|  |    502 | (*                          Main Theorem                              *)
 | 
|  |    503 | (* ------------------------------------------------------------------ *)
 | 
|  |    504 | 
 | 
|  |    505 | lemma compositionality_sch: 
 | 
|  |    506 | "(sch : schedules (A||B)) =  
 | 
|  |    507 |   (Filter (%a. a:act A)$sch : schedules A & 
 | 
|  |    508 |    Filter (%a. a:act B)$sch : schedules B & 
 | 
|  |    509 |    Forall (%x. x:act (A||B)) sch)"
 | 
|  |    510 | apply (simp (no_asm) add: schedules_def has_schedule_def)
 | 
|  |    511 | apply (tactic "safe_tac set_cs")
 | 
|  |    512 | (* ==> *)
 | 
|  |    513 | apply (rule_tac x = "Filter_ex (asig_of A) (ProjA ex) " in bexI)
 | 
|  |    514 | prefer 2
 | 
|  |    515 | apply (simp add: compositionality_ex)
 | 
|  |    516 | apply (simp (no_asm) add: Filter_ex_def ProjA_def lemma_2_1a lemma_2_1b)
 | 
|  |    517 | apply (rule_tac x = "Filter_ex (asig_of B) (ProjB ex) " in bexI)
 | 
|  |    518 | prefer 2
 | 
|  |    519 | apply (simp add: compositionality_ex)
 | 
|  |    520 | apply (simp (no_asm) add: Filter_ex_def ProjB_def lemma_2_1a lemma_2_1b)
 | 
|  |    521 | apply (simp add: executions_def)
 | 
|  |    522 | apply (tactic {* pair_tac "ex" 1 *})
 | 
|  |    523 | apply (erule conjE)
 | 
|  |    524 | apply (simp add: sch_actions_in_AorB)
 | 
|  |    525 | 
 | 
|  |    526 | (* <== *)
 | 
|  |    527 | 
 | 
|  |    528 | (* mkex is exactly the construction of exA||B out of exA, exB, and the oracle sch,
 | 
|  |    529 |    we need here *)
 | 
|  |    530 | apply (rename_tac exA exB)
 | 
|  |    531 | apply (rule_tac x = "mkex A B sch exA exB" in bexI)
 | 
|  |    532 | (* mkex actions are just the oracle *)
 | 
|  |    533 | apply (tactic {* pair_tac "exA" 1 *})
 | 
|  |    534 | apply (tactic {* pair_tac "exB" 1 *})
 | 
|  |    535 | apply (simp add: Mapfst_mkex_is_sch)
 | 
|  |    536 | 
 | 
|  |    537 | (* mkex is an execution -- use compositionality on ex-level *)
 | 
|  |    538 | apply (simp add: compositionality_ex)
 | 
|  |    539 | apply (simp add: stutter_mkex_on_A stutter_mkex_on_B filter_mkex_is_exB filter_mkex_is_exA)
 | 
|  |    540 | apply (tactic {* pair_tac "exA" 1 *})
 | 
|  |    541 | apply (tactic {* pair_tac "exB" 1 *})
 | 
|  |    542 | apply (simp add: mkex_actions_in_AorB)
 | 
|  |    543 | done
 | 
|  |    544 | 
 | 
|  |    545 | 
 | 
|  |    546 | subsection {* COMPOSITIONALITY on SCHEDULE Level -- for Modules *}
 | 
|  |    547 | 
 | 
|  |    548 | lemma compositionality_sch_modules: 
 | 
|  |    549 |   "Scheds (A||B) = par_scheds (Scheds A) (Scheds B)"
 | 
|  |    550 | 
 | 
|  |    551 | apply (unfold Scheds_def par_scheds_def)
 | 
|  |    552 | apply (simp add: asig_of_par)
 | 
|  |    553 | apply (rule set_ext)
 | 
|  |    554 | apply (simp add: compositionality_sch actions_of_par)
 | 
|  |    555 | done
 | 
|  |    556 | 
 | 
|  |    557 | 
 | 
|  |    558 | declare compoex_simps [simp del]
 | 
|  |    559 | declare composch_simps [simp del]
 | 
|  |    560 | 
 | 
|  |    561 | end
 |