src/HOLCF/IOA/meta_theory/Automata.thy
author aspinall
Fri, 30 Sep 2005 18:18:34 +0200
changeset 17740 fc385ce6187d
parent 17233 41eee2e7b465
child 19741 f65265d71426
permissions -rw-r--r--
Add icon for interface.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/meta_theory/Automata.thy
3275
3f53f2c876f4 changes for release 94-8
mueller
parents: 3071
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 12114
diff changeset
     3
    Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     4
*)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     5
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     6
header {* The I/O automata of Lynch and Tuttle in HOLCF *}
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     7
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
theory Automata
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports Asig
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    11
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
defaultsort type
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    13
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    14
types
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    15
  ('a, 's) transition = "'s * 'a * 's"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    16
  ('a, 's) ioa = "'a signature * 's set * ('a,'s)transition set * ('a set set) * ('a set set)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    17
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    18
consts
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    19
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    20
  (* IO automata *)
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    21
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    22
  asig_of        ::"('a,'s)ioa => 'a signature"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    23
  starts_of      ::"('a,'s)ioa => 's set"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    24
  trans_of       ::"('a,'s)ioa => ('a,'s)transition set"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    25
  wfair_of       ::"('a,'s)ioa => ('a set) set"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    26
  sfair_of       ::"('a,'s)ioa => ('a set) set"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    27
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    28
  is_asig_of     ::"('a,'s)ioa => bool"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    29
  is_starts_of   ::"('a,'s)ioa => bool"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    30
  is_trans_of    ::"('a,'s)ioa => bool"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    31
  input_enabled  ::"('a,'s)ioa => bool"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    32
  IOA            ::"('a,'s)ioa => bool"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    33
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    34
  (* constraints for fair IOA *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    35
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    36
  fairIOA        ::"('a,'s)ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    37
  input_resistant::"('a,'s)ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    38
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    39
  (* enabledness of actions and action sets *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    40
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    41
  enabled        ::"('a,'s)ioa => 'a => 's => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    42
  Enabled    ::"('a,'s)ioa => 'a set => 's => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    43
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    44
  (* action set keeps enabled until probably disabled by itself *)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    45
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    46
  en_persistent  :: "('a,'s)ioa => 'a set => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    47
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    48
 (* post_conditions for actions and action sets *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    49
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    50
  was_enabled        ::"('a,'s)ioa => 'a => 's => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    51
  set_was_enabled    ::"('a,'s)ioa => 'a set => 's => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
    52
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    53
  (* reachability and invariants *)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    54
  reachable     :: "('a,'s)ioa => 's set"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    55
  invariant     :: "[('a,'s)ioa, 's=>bool] => bool"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    56
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    57
  (* binary composition of action signatures and automata *)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    58
  asig_comp    ::"['a signature, 'a signature] => 'a signature"
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    59
  compatible   ::"[('a,'s)ioa, ('a,'t)ioa] => bool"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    60
  "||"         ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr 10)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    61
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    62
  (* hiding and restricting *)
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    63
  hide_asig     :: "['a signature, 'a set] => 'a signature"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    64
  "hide"        :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    65
  restrict_asig :: "['a signature, 'a set] => 'a signature"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    66
  restrict      :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    67
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    68
  (* renaming *)
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    69
  rename_set    :: "'a set => ('c => 'a option) => 'c set"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
    70
  rename        :: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    71
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    72
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    73
syntax
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    74
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    75
  "_trans_of"  :: "'s => 'a => ('a,'s)ioa => 's => bool"  ("_ -_--_-> _" [81,81,81,81] 100)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    76
  "reachable"  :: "[('a,'s)ioa, 's] => bool"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    77
  "act"        :: "('a,'s)ioa => 'a set"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    78
  "ext"        :: "('a,'s)ioa => 'a set"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    79
  "int"        :: "('a,'s)ioa => 'a set"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    80
  "inp"        :: "('a,'s)ioa => 'a set"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    81
  "out"        :: "('a,'s)ioa => 'a set"
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
    82
  "local"      :: "('a,'s)ioa => 'a set"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    83
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    84
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 7661
diff changeset
    85
syntax (xsymbols)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    86
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    87
  "_trans_of"  :: "'s => 'a => ('a,'s)ioa => 's => bool"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    88
                  ("_ \<midarrow>_\<midarrow>_\<longrightarrow> _" [81,81,81,81] 100)
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    89
  "op ||"         ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr "\<parallel>" 10)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    90
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    91
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    92
inductive "reachable C"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    93
   intros
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    94
    reachable_0:  "s:(starts_of C) ==> s : reachable C"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    95
    reachable_n:  "[|s:reachable C; (s,a,t):trans_of C|] ==> t:reachable C"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    96
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    97
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    98
translations
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    99
  "s -a--A-> t"   == "(s,a,t):trans_of A"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   100
  "reachable A s" == "s:reachable A"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   101
  "act A"         == "actions (asig_of A)"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   102
  "ext A"         == "externals (asig_of A)"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   103
  "int A"         == "internals (asig_of A)"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   104
  "inp A"         == "inputs (asig_of A)"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   105
  "out A"         == "outputs (asig_of A)"
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
   106
  "local A"       == "locals (asig_of A)"
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
   107
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   108
defs
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   109
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   110
(* --------------------------------- IOA ---------------------------------*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   111
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   112
asig_of_def:   "asig_of == fst"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   113
starts_of_def: "starts_of == (fst o snd)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   114
trans_of_def:  "trans_of == (fst o snd o snd)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   115
wfair_of_def:  "wfair_of == (fst o snd o snd o snd)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   116
sfair_of_def:  "sfair_of == (snd o snd o snd o snd)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   117
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   118
is_asig_of_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   119
  "is_asig_of A == is_asig (asig_of A)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   120
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   121
is_starts_of_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   122
  "is_starts_of A ==  (~ starts_of A = {})"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   123
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   124
is_trans_of_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   125
  "is_trans_of A ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   126
    (!triple. triple:(trans_of A) --> fst(snd(triple)):actions(asig_of A))"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   127
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   128
input_enabled_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   129
  "input_enabled A ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   130
    (!a. (a:inputs(asig_of A)) --> (!s1. ? s2. (s1,a,s2):(trans_of A)))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   131
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   132
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   133
ioa_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   134
  "IOA A == (is_asig_of A    &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   135
             is_starts_of A  &
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   136
             is_trans_of A   &
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   137
             input_enabled A)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   138
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   139
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   140
invariant_def: "invariant A P == (!s. reachable A s --> P(s))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   141
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   142
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   143
(* ------------------------- parallel composition --------------------------*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   144
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   145
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   146
compatible_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   147
  "compatible A B ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   148
  (((out A Int out B) = {}) &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   149
   ((int A Int act B) = {}) &
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
   150
   ((int B Int act A) = {}))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   151
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   152
asig_comp_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   153
  "asig_comp a1 a2 ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   154
     (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)),
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   155
       (outputs(a1) Un outputs(a2)),
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   156
       (internals(a1) Un internals(a2))))"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   157
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   158
par_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   159
  "(A || B) ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   160
      (asig_comp (asig_of A) (asig_of B),
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   161
       {pr. fst(pr):starts_of(A) & snd(pr):starts_of(B)},
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   162
       {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   163
            in (a:act A | a:act B) &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   164
               (if a:act A then
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   165
                  (fst(s),a,fst(t)):trans_of(A)
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   166
                else fst(t) = fst(s))
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   167
               &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   168
               (if a:act B then
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   169
                  (snd(s),a,snd(t)):trans_of(B)
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   170
                else snd(t) = snd(s))},
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   171
        wfair_of A Un wfair_of B,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   172
        sfair_of A Un sfair_of B)"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   173
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   174
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   175
(* ------------------------ hiding -------------------------------------------- *)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   176
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   177
restrict_asig_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   178
  "restrict_asig asig actns ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   179
    (inputs(asig) Int actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   180
     outputs(asig) Int actns,
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   181
     internals(asig) Un (externals(asig) - actns))"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   182
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   183
(* Notice that for wfair_of and sfair_of nothing has to be changed, as
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   184
   changes from the outputs to the internals does not touch the locals as
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   185
   a whole, which is of importance for fairness only *)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   186
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   187
restrict_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   188
  "restrict A actns ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   189
    (restrict_asig (asig_of A) actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   190
     starts_of A,
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   191
     trans_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   192
     wfair_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   193
     sfair_of A)"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   194
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   195
hide_asig_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   196
  "hide_asig asig actns ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   197
    (inputs(asig) - actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   198
     outputs(asig) - actns,
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   199
     internals(asig) Un actns)"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   200
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   201
hide_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   202
  "hide A actns ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   203
    (hide_asig (asig_of A) actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   204
     starts_of A,
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   205
     trans_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   206
     wfair_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   207
     sfair_of A)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   208
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   209
(* ------------------------- renaming ------------------------------------------- *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   210
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   211
rename_set_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   212
  "rename_set A ren == {b. ? x. Some x = ren b & x : A}"
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   213
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   214
rename_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   215
"rename ioa ren ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   216
  ((rename_set (inp ioa) ren,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   217
    rename_set (out ioa) ren,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   218
    rename_set (int ioa) ren),
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   219
   starts_of ioa,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   220
   {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   221
        in
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   222
        ? x. Some(x) = ren(a) & (s,x,t):trans_of ioa},
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   223
   {rename_set s ren | s. s: wfair_of ioa},
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   224
   {rename_set s ren | s. s: sfair_of ioa})"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   225
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   226
(* ------------------------- fairness ----------------------------- *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   227
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   228
fairIOA_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   229
  "fairIOA A == (! S : wfair_of A. S<= local A) &
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   230
                (! S : sfair_of A. S<= local A)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   231
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   232
input_resistant_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   233
  "input_resistant A == ! W : sfair_of A. ! s a t.
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   234
                        reachable A s & reachable A t & a:inp A &
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   235
                        Enabled A W s & s -a--A-> t
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   236
                        --> Enabled A W t"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   237
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   238
enabled_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   239
  "enabled A a s == ? t. s-a--A-> t"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   240
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   241
Enabled_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   242
  "Enabled A W s == ? w:W. enabled A w s"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   243
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   244
en_persistent_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   245
  "en_persistent A W == ! s a t. Enabled A W s &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   246
                                 a ~:W &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   247
                                 s -a--A-> t
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   248
                                 --> Enabled A W t"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   249
was_enabled_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   250
  "was_enabled A a t == ? s. s-a--A-> t"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   251
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   252
set_was_enabled_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   253
  "set_was_enabled A W t == ? w:W. was_enabled A w t"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   254
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   255
ML {* use_legacy_bindings (the_context ()) *}
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   256
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   257
end