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