| author | isatest | 
| Thu, 22 Sep 2005 00:30:31 +0200 | |
| changeset 17578 | e07af5fad73f | 
| parent 17288 | aa3833fb7bee | 
| child 19801 | b2af2549efd1 | 
| permissions | -rw-r--r-- | 
| 4530 | 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 | 7 | header {* The I/O automata of Lynch and Tuttle *}
 | 
| 8 | ||
| 9 | theory IOA | |
| 10 | imports Asig | |
| 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 | 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 | 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 | 67 | state_trans_def: | 
| 68 | "state_trans asig R == | |
| 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 | 73 | asig_of_def: "asig_of == fst" | 
| 74 | starts_of_def: "starts_of == (fst o snd)" | |
| 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 | 77 | ioa_def: | 
| 78 | "IOA(ioa) == (is_asig(asig_of(ioa)) & | |
| 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 | 87 | is_execution_fragment_def: | 
| 88 | "is_execution_fragment A ex == | |
| 89 | let act = fst(ex); state = snd(ex) | |
| 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 | 94 | executions_def: | 
| 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 | 98 | |
| 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 | 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 | 107 | filter_oseq_def: | 
| 108 | "filter_oseq p s == | |
| 109 | (%i. case s(i) | |
| 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 | 114 | mk_trace_def: | 
| 3842 | 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 | 119 | has_trace_def: | 
| 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 | 123 | normal_form_def: | 
| 124 | "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) & | |
| 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 | 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 | 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 | 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 | 137 | |
| 138 | compat_asigs_def: | |
| 139 | "compat_asigs a1 a2 == | |
| 140 |    (((outputs(a1) Int outputs(a2)) = {}) &
 | |
| 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 | 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 | 149 | asig_comp_def: | 
| 150 | "asig_comp a1 a2 == | |
| 151 | (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)), | |
| 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 | 156 | par_def: | 
| 157 | "(ioa1 || ioa2) == | |
| 158 | (asig_comp (asig_of ioa1) (asig_of ioa2), | |
| 159 |         {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)},
 | |
| 160 |         {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
 | |
| 161 | in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) & | |
| 162 | (if a:actions(asig_of(ioa1)) then | |
| 163 | (fst(s),a,fst(t)):trans_of(ioa1) | |
| 164 | else fst(t) = fst(s)) | |
| 165 | & | |
| 166 | (if a:actions(asig_of(ioa2)) then | |
| 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 | 171 | restrict_asig_def: | 
| 172 | "restrict_asig asig actns == | |
| 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 | 177 | restrict_def: | 
| 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 | 182 | ioa_implements_def: | 
| 183 | "ioa_implements ioa1 ioa2 == | |
| 184 | ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) & | |
| 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 | 187 | |
| 188 | rename_def: | |
| 189 | "rename ioa ren == | |
| 190 |   (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))},
 | |
| 191 |     {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))},
 | |
| 192 |     {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}),
 | |
| 193 | starts_of(ioa) , | |
| 194 |    {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
 | |
| 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 | 198 | ML {* use_legacy_bindings (the_context ()) *}
 | 
| 199 | ||
| 200 | end |