src/HOL/IOA/IOA.thy
author wenzelm
Tue, 06 Sep 2005 19:03:39 +0200
changeset 17288 aa3833fb7bee
parent 11482 ec2c382ff4f0
child 19801 b2af2549efd1
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4530
ac1821645636 corrected Title
oheimb
parents: 3842
diff changeset
     1
(*  Title:      HOL/IOA/IOA.thy
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     2
    ID:         $Id$
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     4
    Copyright   1994  TU Muenchen
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     5
*)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     6
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
     7
header {* The I/O automata of Lynch and Tuttle *}
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
     8
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
     9
theory IOA
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    10
imports Asig
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    11
begin
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    12
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    13
types
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    14
   'a seq            =   "nat => 'a"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    15
   'a oseq           =   "nat => 'a option"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    16
   ('a,'b)execution  =   "'a oseq * 'b seq"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    17
   ('a,'s)transition =   "('s * 'a * 's)"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    18
   ('a,'s)ioa        =   "'a signature * 's set * ('a,'s)transition set"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    19
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    20
consts
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    21
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    22
  (* IO automata *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    23
  state_trans::"['action signature, ('action,'state)transition set] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    24
  asig_of    ::"('action,'state)ioa => 'action signature"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    25
  starts_of  ::"('action,'state)ioa => 'state set"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    26
  trans_of   ::"('action,'state)ioa => ('action,'state)transition set"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    27
  IOA        ::"('action,'state)ioa => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    28
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    29
  (* Executions, schedules, and traces *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    30
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    31
  is_execution_fragment ::"[('action,'state)ioa, ('action,'state)execution] => bool"
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    32
  has_execution ::"[('action,'state)ioa, ('action,'state)execution] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    33
  executions    :: "('action,'state)ioa => ('action,'state)execution set"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    34
  mk_trace  :: "[('action,'state)ioa, 'action oseq] => 'action oseq"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    35
  reachable     :: "[('action,'state)ioa, 'state] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    36
  invariant     :: "[('action,'state)ioa, 'state=>bool] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    37
  has_trace :: "[('action,'state)ioa, 'action oseq] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    38
  traces    :: "('action,'state)ioa => 'action oseq set"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    39
  NF            :: "'a oseq => 'a oseq"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    40
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    41
  (* Composition of action signatures and automata *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    42
  compatible_asigs ::"('a => 'action signature) => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    43
  asig_composition ::"('a => 'action signature) => 'action signature"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    44
  compatible_ioas  ::"('a => ('action,'state)ioa) => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    45
  ioa_composition  ::"('a => ('action, 'state)ioa) =>('action,'a => 'state)ioa"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    46
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    47
  (* binary composition of action signatures and automata *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    48
  compat_asigs ::"['action signature, 'action signature] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    49
  asig_comp    ::"['action signature, 'action signature] => 'action signature"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    50
  compat_ioas  ::"[('action,'s)ioa, ('action,'t)ioa] => bool"
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    51
  par         ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr "||" 10)
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    52
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    53
  (* Filtering and hiding *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    54
  filter_oseq  :: "('a => bool) => 'a oseq => 'a oseq"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    55
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    56
  restrict_asig :: "['a signature, 'a set] => 'a signature"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    57
  restrict      :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    58
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    59
  (* Notions of correctness *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    60
  ioa_implements :: "[('action,'state1)ioa, ('action,'state2)ioa] => bool"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    61
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    62
  (* Instantiation of abstract IOA by concrete actions *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    63
  rename:: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    64
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    65
defs
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    66
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    67
state_trans_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    68
  "state_trans asig R ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    69
     (!triple. triple:R --> fst(snd(triple)):actions(asig)) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    70
     (!a. (a:inputs(asig)) --> (!s1. ? s2. (s1,a,s2):R))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    71
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    72
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    73
asig_of_def:   "asig_of == fst"
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    74
starts_of_def: "starts_of == (fst o snd)"
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    75
trans_of_def:  "trans_of == (snd o snd)"
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    76
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    77
ioa_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    78
  "IOA(ioa) == (is_asig(asig_of(ioa))      &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    79
                (~ starts_of(ioa) = {})    &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    80
                state_trans (asig_of ioa) (trans_of ioa))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    81
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    82
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    83
(* An execution fragment is modelled with a pair of sequences:
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    84
 * the first is the action options, the second the state sequence.
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    85
 * Finite executions have None actions from some point on.
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    86
 *******)
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    87
is_execution_fragment_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    88
  "is_execution_fragment A ex ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    89
     let act = fst(ex); state = snd(ex)
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    90
     in !n a. (act(n)=None --> state(Suc(n)) = state(n)) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    91
              (act(n)=Some(a) --> (state(n),a,state(Suc(n))):trans_of(A))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    92
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    93
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    94
executions_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    95
  "executions(ioa) == {e. snd e 0:starts_of(ioa) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    96
                        is_execution_fragment ioa e}"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    97
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    98
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
    99
reachable_def:
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   100
  "reachable ioa s == (? ex:executions(ioa). ? n. (snd ex n) = s)"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   101
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   102
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   103
invariant_def: "invariant A P == (!s. reachable A s --> P(s))"
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   104
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   105
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   106
(* Restrict the trace to those members of the set s *)
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   107
filter_oseq_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   108
  "filter_oseq p s ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   109
   (%i. case s(i)
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   110
         of None => None
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   111
          | Some(x) => if p x then Some x else None)"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   112
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   113
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   114
mk_trace_def:
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3078
diff changeset
   115
  "mk_trace(ioa) == filter_oseq(%a. a:externals(asig_of(ioa)))"
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   116
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   117
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   118
(* Does an ioa have an execution with the given trace *)
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   119
has_trace_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   120
  "has_trace ioa b ==
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   121
     (? ex:executions(ioa). b = mk_trace ioa (fst ex))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   122
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   123
normal_form_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   124
  "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   125
                    (!j. j ~: range(f) --> nf(j)= None) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   126
                    (!i. nf(i)=None --> (nf (Suc i)) = None)   "
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   127
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   128
(* All the traces of an ioa *)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   129
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   130
  traces_def:
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   131
  "traces(ioa) == {trace. ? tr. trace=NF(tr) & has_trace ioa tr}"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   132
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   133
(*
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   134
  traces_def:
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   135
  "traces(ioa) == {tr. has_trace ioa tr}"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   136
*)
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   137
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   138
compat_asigs_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   139
  "compat_asigs a1 a2 ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   140
   (((outputs(a1) Int outputs(a2)) = {}) &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   141
    ((internals(a1) Int actions(a2)) = {}) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   142
    ((internals(a2) Int actions(a1)) = {}))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   143
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   144
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   145
compat_ioas_def:
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   146
  "compat_ioas ioa1 ioa2 == compat_asigs (asig_of(ioa1)) (asig_of(ioa2))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   147
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   148
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   149
asig_comp_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   150
  "asig_comp a1 a2 ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   151
      (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)),
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   152
        (outputs(a1) Un outputs(a2)),
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   153
        (internals(a1) Un internals(a2))))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   154
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   155
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   156
par_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   157
  "(ioa1 || ioa2) ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   158
       (asig_comp (asig_of ioa1) (asig_of ioa2),
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   159
        {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)},
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   160
        {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   161
             in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   162
                (if a:actions(asig_of(ioa1)) then
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   163
                   (fst(s),a,fst(t)):trans_of(ioa1)
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   164
                 else fst(t) = fst(s))
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   165
                &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   166
                (if a:actions(asig_of(ioa2)) then
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   167
                   (snd(s),a,snd(t)):trans_of(ioa2)
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   168
                 else snd(t) = snd(s))})"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   169
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   170
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   171
restrict_asig_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   172
  "restrict_asig asig actns ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   173
    (inputs(asig) Int actns, outputs(asig) Int actns,
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   174
     internals(asig) Un (externals(asig) - actns))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   175
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   176
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   177
restrict_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   178
  "restrict ioa actns ==
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   179
    (restrict_asig (asig_of ioa) actns, starts_of(ioa), trans_of(ioa))"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   180
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   181
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   182
ioa_implements_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   183
  "ioa_implements ioa1 ioa2 ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   184
  ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) &
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   185
     (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) &
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   186
      traces(ioa1) <= traces(ioa2))"
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   187
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   188
rename_def:
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   189
"rename ioa ren ==
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   190
  (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))},
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   191
    {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))},
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   192
    {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}),
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   193
              starts_of(ioa)   ,
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   194
   {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   195
        in
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   196
        ? x. Some(x) = ren(a) & (s,x,t):trans_of(ioa)})"
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
   197
17288
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   198
ML {* use_legacy_bindings (the_context ()) *}
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   199
aa3833fb7bee converted to Isar theory format;
wenzelm
parents: 11482
diff changeset
   200
end