| 3071 |      1 | (*  Title:      HOLCF/IOA/meta_theory/CompoTraces.thy
 | 
| 3275 |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 3071 |      4 | 
 | 
|  |      5 | Compositionality on Trace level.
 | 
|  |      6 | *) 
 | 
|  |      7 | 
 | 
|  |      8 | CompoTraces = CompoScheds + ShortExecutions +
 | 
|  |      9 |  
 | 
|  |     10 | 
 | 
|  |     11 | consts  
 | 
|  |     12 | 
 | 
| 3521 |     13 |  mksch      ::"('a,'s)ioa => ('a,'t)ioa => 'a Seq -> 'a Seq -> 'a Seq -> 'a Seq" 
 | 
|  |     14 |  par_traces ::"['a trace_module,'a trace_module] => 'a trace_module"
 | 
| 3071 |     15 | 
 | 
|  |     16 | defs
 | 
|  |     17 | 
 | 
|  |     18 | mksch_def
 | 
| 10835 |     19 |   "mksch A B == (fix$(LAM h tr schA schB. case tr of 
 | 
| 3071 |     20 |        nil => nil
 | 
|  |     21 |     | x##xs => 
 | 
|  |     22 |       (case x of 
 | 
| 12028 |     23 |         UU => UU
 | 
| 3071 |     24 |       | Def y => 
 | 
|  |     25 |          (if y:act A then 
 | 
|  |     26 |              (if y:act B then 
 | 
| 10835 |     27 |                    ((Takewhile (%a. a:int A)$schA)
 | 
|  |     28 |                       @@ (Takewhile (%a. a:int B)$schB)
 | 
|  |     29 |                            @@ (y>>(h$xs
 | 
|  |     30 |                                     $(TL$(Dropwhile (%a. a:int A)$schA))
 | 
|  |     31 |                                     $(TL$(Dropwhile (%a. a:int B)$schB))
 | 
| 3071 |     32 |                     )))
 | 
|  |     33 |               else
 | 
| 10835 |     34 |                  ((Takewhile (%a. a:int A)$schA)
 | 
|  |     35 |                   @@ (y>>(h$xs
 | 
|  |     36 |                            $(TL$(Dropwhile (%a. a:int A)$schA))
 | 
|  |     37 |                            $schB)))
 | 
| 3071 |     38 |               )
 | 
|  |     39 |           else 
 | 
|  |     40 |              (if y:act B then 
 | 
| 10835 |     41 |                  ((Takewhile (%a. a:int B)$schB)
 | 
|  |     42 |                      @@ (y>>(h$xs
 | 
|  |     43 |                               $schA
 | 
|  |     44 |                               $(TL$(Dropwhile (%a. a:int B)$schB))
 | 
| 3071 |     45 |                               )))
 | 
|  |     46 |              else
 | 
|  |     47 |                UU
 | 
|  |     48 |              )
 | 
|  |     49 |          )
 | 
|  |     50 |        )))"
 | 
|  |     51 | 
 | 
|  |     52 | 
 | 
| 3521 |     53 | par_traces_def
 | 
|  |     54 |   "par_traces TracesA TracesB == 
 | 
|  |     55 |        let trA = fst TracesA; sigA = snd TracesA; 
 | 
|  |     56 |            trB = fst TracesB; sigB = snd TracesB       
 | 
|  |     57 |        in
 | 
| 10835 |     58 |        (    {tr. Filter (%a. a:actions sigA)$tr : trA}
 | 
|  |     59 |         Int {tr. Filter (%a. a:actions sigB)$tr : trB}
 | 
| 3521 |     60 |         Int {tr. Forall (%x. x:(externals sigA Un externals sigB)) tr},
 | 
|  |     61 |         asig_comp sigA sigB)"
 | 
|  |     62 | 
 | 
| 3071 |     63 | rules
 | 
|  |     64 | 
 | 
|  |     65 | 
 | 
|  |     66 | finiteR_mksch
 | 
| 10835 |     67 |   "Finite (mksch A B$tr$x$y) --> Finite tr"
 | 
| 3071 |     68 | 
 | 
|  |     69 | 
 | 
|  |     70 | 
 | 
|  |     71 | end
 |