| 
4559
 | 
     1  | 
(*  Title:      HOLCF/IOA/meta_theory/LiveIOA.thy
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
12218
 | 
     3  | 
    Author:     Olaf Müller
  | 
| 
17233
 | 
     4  | 
*)
  | 
| 
4559
 | 
     5  | 
  | 
| 
17233
 | 
     6  | 
header {* Live I/O automata -- specified by temproal formulas *}
 | 
| 
4559
 | 
     7  | 
  | 
| 
17233
 | 
     8  | 
theory LiveIOA
  | 
| 
 | 
     9  | 
imports TLS
  | 
| 
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
defaultsort type
  | 
| 
4559
 | 
    13  | 
  | 
| 
 | 
    14  | 
types
  | 
| 
17233
 | 
    15  | 
  ('a, 's) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp"
 | 
| 
4559
 | 
    16  | 
  | 
| 
 | 
    17  | 
consts
  | 
| 
 | 
    18  | 
  | 
| 
 | 
    19  | 
validLIOA   :: "('a,'s)live_ioa => ('a,'s)ioa_temp  => bool"
 | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
WF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
 | 
| 
 | 
    22  | 
SF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
 | 
| 
 | 
    23  | 
  | 
| 
 | 
    24  | 
liveexecutions    :: "('a,'s)live_ioa => ('a,'s)execution set"
 | 
| 
 | 
    25  | 
livetraces        :: "('a,'s)live_ioa => 'a trace set"
 | 
| 
 | 
    26  | 
live_implements   :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
 | 
| 
 | 
    27  | 
is_live_ref_map   :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
 | 
| 
 | 
    28  | 
  | 
| 
17233
 | 
    29  | 
  | 
| 
4559
 | 
    30  | 
defs
  | 
| 
 | 
    31  | 
  | 
| 
17233
 | 
    32  | 
validLIOA_def:
  | 
| 
4559
 | 
    33  | 
  "validLIOA AL P == validIOA (fst AL) ((snd AL) .--> P)"
  | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
  | 
| 
17233
 | 
    36  | 
WF_def:
  | 
| 
4559
 | 
    37  | 
  "WF A acts ==  <> [] <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
  | 
| 
 | 
    38  | 
  | 
| 
17233
 | 
    39  | 
SF_def:
  | 
| 
4559
 | 
    40  | 
  "SF A acts ==  [] <> <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
  | 
| 
17233
 | 
    41  | 
  | 
| 
4559
 | 
    42  | 
  | 
| 
17233
 | 
    43  | 
liveexecutions_def:
  | 
| 
4559
 | 
    44  | 
   "liveexecutions AP == {exec. exec : executions (fst AP) & (exec |== (snd AP))}"
 | 
| 
 | 
    45  | 
  | 
| 
17233
 | 
    46  | 
livetraces_def:
  | 
| 
10835
 | 
    47  | 
  "livetraces AP == {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}"
 | 
| 
4559
 | 
    48  | 
  | 
| 
17233
 | 
    49  | 
live_implements_def:
  | 
| 
 | 
    50  | 
  "live_implements CL AM == (inp (fst CL) = inp (fst AM)) &
  | 
| 
4559
 | 
    51  | 
                            (out (fst CL) = out (fst AM)) &
  | 
| 
 | 
    52  | 
                            livetraces CL <= livetraces AM"
  | 
| 
 | 
    53  | 
  | 
| 
17233
 | 
    54  | 
is_live_ref_map_def:
  | 
| 
 | 
    55  | 
   "is_live_ref_map f CL AM ==
  | 
| 
 | 
    56  | 
            is_ref_map f (fst CL ) (fst AM) &
  | 
| 
 | 
    57  | 
            (! exec : executions (fst CL). (exec |== (snd CL)) -->
  | 
| 
4559
 | 
    58  | 
                                           ((corresp_ex (fst AM) f exec) |== (snd AM)))"
  | 
| 
 | 
    59  | 
  | 
| 
17233
 | 
    60  | 
ML {* use_legacy_bindings (the_context ()) *}
 | 
| 
4559
 | 
    61  | 
  | 
| 
17233
 | 
    62  | 
end
  |