src/HOL/HOLCF/IOA/Automata.thy
author haftmann
Mon, 09 Oct 2017 19:10:52 +0200
changeset 66840 0d689d71dbdc
parent 62116 bc178c0fe1a1
child 69597 ff784d5a5bfb
permissions -rw-r--r--
canonical multiplicative euclidean size
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62005
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/Automata.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     3
*)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 61999
diff changeset
     5
section \<open>The I/O automata of Lynch and Tuttle in HOLCF\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory Automata
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Asig
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 36176
diff changeset
    11
default_sort type
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
type_synonym ('a, 's) transition = "'s \<times> 'a \<times> 's"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    14
type_synonym ('a, 's) ioa =
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    15
  "'a signature \<times> 's set \<times> ('a, 's)transition set \<times> 'a set set \<times> 'a set set"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    16
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    17
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    18
subsection \<open>IO automata\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    19
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    20
definition asig_of :: "('a, 's) ioa \<Rightarrow> 'a signature"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    21
  where "asig_of = fst"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    22
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    23
definition starts_of :: "('a, 's) ioa \<Rightarrow> 's set"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    24
  where "starts_of = fst \<circ> snd"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    25
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    26
definition trans_of :: "('a, 's) ioa \<Rightarrow> ('a, 's) transition set"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    27
  where "trans_of = fst \<circ> snd \<circ> snd"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    28
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    29
abbreviation trans_of_syn  ("_ \<midarrow>_\<midarrow>_\<rightarrow> _" [81, 81, 81, 81] 100)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    30
  where "s \<midarrow>a\<midarrow>A\<rightarrow> t \<equiv> (s, a, t) \<in> trans_of A"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    31
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    32
definition wfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
  where "wfair_of = fst \<circ> snd \<circ> snd \<circ> snd"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    34
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    35
definition sfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    36
  where "sfair_of = snd \<circ> snd \<circ> snd \<circ> snd"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    37
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    38
definition is_asig_of :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    39
  where "is_asig_of A = is_asig (asig_of A)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    40
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    41
definition is_starts_of :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    42
  where "is_starts_of A \<longleftrightarrow> starts_of A \<noteq> {}"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    43
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    44
definition is_trans_of :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    45
  where "is_trans_of A \<longleftrightarrow>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    46
    (\<forall>triple. triple \<in> trans_of A \<longrightarrow> fst (snd triple) \<in> actions (asig_of A))"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    47
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    48
definition input_enabled :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    49
  where "input_enabled A \<longleftrightarrow>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    50
    (\<forall>a. a \<in> inputs (asig_of A) \<longrightarrow> (\<forall>s1. \<exists>s2. (s1, a, s2) \<in> trans_of A))"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    51
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    52
definition IOA :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    53
  where "IOA A \<longleftrightarrow>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    54
    is_asig_of A \<and>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    55
    is_starts_of A \<and>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    56
    is_trans_of A \<and>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    57
    input_enabled A"
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    58
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    59
abbreviation "act A \<equiv> actions (asig_of A)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    60
abbreviation "ext A \<equiv> externals (asig_of A)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
abbreviation int where "int A \<equiv> internals (asig_of A)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    62
abbreviation "inp A \<equiv> inputs (asig_of A)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    63
abbreviation "out A \<equiv> outputs (asig_of A)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    64
abbreviation "local A \<equiv> locals (asig_of A)"
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
    65
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    66
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    67
text \<open>invariants\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    68
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    69
inductive reachable :: "('a, 's) ioa \<Rightarrow> 's \<Rightarrow> bool" for C :: "('a, 's) ioa"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    70
where
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    71
  reachable_0:  "s \<in> starts_of C \<Longrightarrow> reachable C s"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    72
| reachable_n:  "reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> reachable C t"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    73
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    74
definition invariant :: "[('a, 's) ioa, 's \<Rightarrow> bool] \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    75
  where "invariant A P \<longleftrightarrow> (\<forall>s. reachable A s \<longrightarrow> P s)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    76
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    77
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    78
subsection \<open>Parallel composition\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    79
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    80
subsubsection \<open>Binary composition of action signatures and automata\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    81
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
definition compatible :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> bool"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
  where "compatible A B \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    84
    out A \<inter> out B = {} \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    85
    int A \<inter> act B = {} \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    86
    int B \<inter> act A = {}"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    87
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    88
definition asig_comp :: "'a signature \<Rightarrow> 'a signature \<Rightarrow> 'a signature"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    89
  where "asig_comp a1 a2 =
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    90
     (((inputs a1 \<union> inputs a2) - (outputs a1 \<union> outputs a2),
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    91
       (outputs a1 \<union> outputs a2),
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    92
       (internals a1 \<union> internals a2)))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    93
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    94
definition par :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> ('a, 's * 't) ioa"  (infixr "\<parallel>" 10)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    95
  where "(A \<parallel> B) =
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    96
    (asig_comp (asig_of A) (asig_of B),
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    97
     {pr. fst pr \<in> starts_of A \<and> snd pr \<in> starts_of B},
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    98
     {tr.
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    99
        let
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   100
          s = fst tr;
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   101
          a = fst (snd tr);
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   102
          t = snd (snd tr)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   103
        in
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   104
          (a \<in> act A \<or> a \<in> act B) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   105
          (if a \<in> act A then (fst s, a, fst t) \<in> trans_of A
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   106
           else fst t = fst s) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   107
          (if a \<in> act B then (snd s, a, snd t) \<in> trans_of B
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   108
           else snd t = snd s)},
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   109
      wfair_of A \<union> wfair_of B,
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   110
      sfair_of A \<union> sfair_of B)"
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   111
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   112
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   113
subsection \<open>Hiding\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   114
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   115
subsubsection \<open>Hiding and restricting\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   116
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   117
definition restrict_asig :: "'a signature \<Rightarrow> 'a set \<Rightarrow> 'a signature"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   118
  where "restrict_asig asig actns =
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   119
    (inputs asig \<inter> actns,
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   120
     outputs asig \<inter> actns,
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   121
     internals asig \<union> (externals asig - actns))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   122
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   123
text \<open>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   124
  Notice that for \<open>wfair_of\<close> and \<open>sfair_of\<close> nothing has to be changed, as
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   125
  changes from the outputs to the internals does not touch the locals as a
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   126
  whole, which is of importance for fairness only.
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   127
\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   128
definition restrict :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) ioa"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   129
  where "restrict A actns =
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   130
    (restrict_asig (asig_of A) actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   131
     starts_of A,
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   132
     trans_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   133
     wfair_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   134
     sfair_of A)"
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   135
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   136
definition hide_asig :: "'a signature \<Rightarrow> 'a set \<Rightarrow> 'a signature"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   137
  where "hide_asig asig actns =
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   138
    (inputs asig - actns,
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   139
     outputs asig - actns,
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   140
     internals asig \<union> actns)"
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   141
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   142
definition hide :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) ioa"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   143
  where "hide A actns =
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   144
    (hide_asig (asig_of A) actns,
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   145
     starts_of A,
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   146
     trans_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   147
     wfair_of A,
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   148
     sfair_of A)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   149
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   150
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   151
subsection \<open>Renaming\<close>
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   152
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   153
definition rename_set :: "'a set \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> 'c set"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   154
  where "rename_set A ren = {b. \<exists>x. Some x = ren b \<and> x \<in> A}"
3521
bdc51b4c6050 changes needed for adding fairness
mueller
parents: 3433
diff changeset
   155
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   156
definition rename :: "('a, 'b) ioa \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> ('c, 'b) ioa"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   157
  where "rename ioa ren =
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   158
    ((rename_set (inp ioa) ren,
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   159
      rename_set (out ioa) ren,
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   160
      rename_set (int ioa) ren),
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   161
     starts_of ioa,
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   162
     {tr.
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   163
        let
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   164
          s = fst tr;
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   165
          a = fst (snd tr);
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   166
          t = snd (snd tr)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   167
        in \<exists>x. Some x = ren a \<and> s \<midarrow>x\<midarrow>ioa\<rightarrow> t},
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   168
     {rename_set s ren | s. s \<in> wfair_of ioa},
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   169
     {rename_set s ren | s. s \<in> sfair_of ioa})"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   170
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   171
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   172
subsection \<open>Fairness\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   173
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   174
subsubsection \<open>Enabledness of actions and action sets\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   175
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   176
definition enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   177
  where "enabled A a s \<longleftrightarrow> (\<exists>t. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   178
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   179
definition Enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   180
  where "Enabled A W s \<longleftrightarrow> (\<exists>w \<in> W. enabled A w s)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   181
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   182
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   183
text \<open>Action set keeps enabled until probably disabled by itself.\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   184
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   185
definition en_persistent :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> bool"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   186
  where "en_persistent A W \<longleftrightarrow>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   187
    (\<forall>s a t. Enabled A W s \<and> a \<notin> W \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   188
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   189
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   190
text \<open>Post conditions for actions and action sets.\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   191
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   192
definition was_enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   193
  where "was_enabled A a t \<longleftrightarrow> (\<exists>s. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   194
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   195
definition set_was_enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   196
  where "set_was_enabled A W t \<longleftrightarrow> (\<exists>w \<in> W. was_enabled A w t)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   197
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   198
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   199
text \<open>Constraints for fair IOA.\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   200
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   201
definition fairIOA :: "('a, 's) ioa \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   202
  where "fairIOA A \<longleftrightarrow> (\<forall>S \<in> wfair_of A. S \<subseteq> local A) \<and> (\<forall>S \<in> sfair_of A. S \<subseteq> local A)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   203
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   204
definition input_resistant :: "('a, 's) ioa \<Rightarrow> bool"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   205
  where "input_resistant A \<longleftrightarrow>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   206
    (\<forall>W \<in> sfair_of A. \<forall>s a t.
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   207
      reachable A s \<and> reachable A t \<and> a \<in> inp A \<and>
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   208
      Enabled A W s \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3521
diff changeset
   209
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   210
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   211
declare split_paired_Ex [simp del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   212
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   213
lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wfair_of_def sfair_of_def
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   214
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   215
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   216
subsection "\<open>asig_of\<close>, \<open>starts_of\<close>, \<open>trans_of\<close>"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   217
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   218
lemma ioa_triple_proj:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   219
  "asig_of (x, y, z, w, s) = x \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   220
   starts_of (x, y, z, w, s) = y \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   221
   trans_of (x, y, z, w, s) = z \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   222
   wfair_of (x, y, z, w, s) = w \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   223
   sfair_of (x, y, z, w, s) = s"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   224
  by (simp add: ioa_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   225
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   226
lemma trans_in_actions: "is_trans_of A \<Longrightarrow> s1 \<midarrow>a\<midarrow>A\<rightarrow> s2 \<Longrightarrow> a \<in> act A"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   227
  by (auto simp add: is_trans_of_def actions_def is_asig_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   228
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   229
lemma starts_of_par: "starts_of (A \<parallel> B) = {p. fst p \<in> starts_of A \<and> snd p \<in> starts_of B}"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   230
  by (simp add: par_def ioa_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   231
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   232
lemma trans_of_par:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   233
  "trans_of(A \<parallel> B) =
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   234
    {tr.
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   235
      let
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   236
        s = fst tr;
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   237
        a = fst (snd tr);
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   238
        t = snd (snd tr)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   239
      in
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   240
        (a \<in> act A \<or> a \<in> act B) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   241
        (if a \<in> act A then (fst s, a, fst t) \<in> trans_of A
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   242
         else fst t = fst s) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   243
        (if a \<in> act B then (snd s, a, snd t) \<in> trans_of B
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   244
         else snd t = snd s)}"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   245
  by (simp add: par_def ioa_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   246
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   247
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   248
subsection \<open>\<open>actions\<close> and \<open>par\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   249
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   250
lemma actions_asig_comp: "actions (asig_comp a b) = actions a \<union> actions b"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   251
  by (auto simp add: actions_def asig_comp_def asig_projections)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   252
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   253
lemma asig_of_par: "asig_of(A \<parallel> B) = asig_comp (asig_of A) (asig_of B)"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   254
  by (simp add: par_def ioa_projections)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   255
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   256
lemma externals_of_par: "ext (A1 \<parallel> A2) = ext A1 \<union> ext A2"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   257
  by (auto simp add: externals_def asig_of_par asig_comp_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   258
    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   259
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   260
lemma actions_of_par: "act (A1 \<parallel> A2) = act A1 \<union> act A2"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   261
  by (auto simp add: actions_def asig_of_par asig_comp_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   262
    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   263
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   264
lemma inputs_of_par: "inp (A1 \<parallel> A2) = (inp A1 \<union> inp A2) - (out A1 \<union> out A2)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   265
  by (simp add: actions_def asig_of_par asig_comp_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   266
    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   267
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   268
lemma outputs_of_par: "out (A1 \<parallel> A2) = out A1 \<union> out A2"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   269
  by (simp add: actions_def asig_of_par asig_comp_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   270
    asig_outputs_def Un_def set_diff_eq)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   271
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   272
lemma internals_of_par: "int (A1 \<parallel> A2) = int A1 \<union> int A2"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   273
  by (simp add: actions_def asig_of_par asig_comp_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   274
    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   275
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   276
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   277
subsection \<open>Actions and compatibility\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   278
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   279
lemma compat_commute: "compatible A B = compatible B A"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   280
  by (auto simp add: compatible_def Int_commute)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   281
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   282
lemma ext1_is_not_int2: "compatible A1 A2 \<Longrightarrow> a \<in> ext A1 \<Longrightarrow> a \<notin> int A2"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   283
  by (auto simp add: externals_def actions_def compatible_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   284
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   285
(*just commuting the previous one: better commute compatible*)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   286
lemma ext2_is_not_int1: "compatible A2 A1 \<Longrightarrow> a \<in> ext A1 \<Longrightarrow> a \<notin> int A2"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   287
  by (auto simp add: externals_def actions_def compatible_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   288
45606
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 44066
diff changeset
   289
lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act]
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 44066
diff changeset
   290
lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   291
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   292
lemma intA_is_not_extB: "compatible A B \<Longrightarrow> x \<in> int A \<Longrightarrow> x \<notin> ext B"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   293
  by (auto simp add: externals_def actions_def compatible_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   294
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   295
lemma intA_is_not_actB: "compatible A B \<Longrightarrow> a \<in> int A \<Longrightarrow> a \<notin> act B"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   296
  by (auto simp add: externals_def actions_def compatible_def is_asig_def asig_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   297
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   298
(*the only one that needs disjointness of outputs and of internals and _all_ acts*)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   299
lemma outAactB_is_inpB: "compatible A B \<Longrightarrow> a \<in> out A \<Longrightarrow> a \<in> act B \<Longrightarrow> a \<in> inp B"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   300
  by (auto simp add: asig_outputs_def asig_internals_def actions_def asig_inputs_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   301
      compatible_def is_asig_def asig_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   302
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   303
(*needed for propagation of input_enabledness from A, B to A \<parallel> B*)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   304
lemma inpAAactB_is_inpBoroutB:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   305
  "compatible A B \<Longrightarrow> a \<in> inp A \<Longrightarrow> a \<in> act B \<Longrightarrow> a \<in> inp B \<or> a \<in> out B"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   306
  by (auto simp add: asig_outputs_def asig_internals_def actions_def asig_inputs_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   307
      compatible_def is_asig_def asig_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   308
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   309
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   310
subsection \<open>Input enabledness and par\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   311
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   312
(*ugly case distinctions. Heart of proof:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   313
    1. inpAAactB_is_inpBoroutB ie. internals are really hidden.
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   314
    2. inputs_of_par: outputs are no longer inputs of par. This is important here.*)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   315
lemma input_enabled_par:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   316
  "compatible A B \<Longrightarrow> input_enabled A \<Longrightarrow> input_enabled B \<Longrightarrow> input_enabled (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   317
  apply (unfold input_enabled_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   318
  apply (simp add: Let_def inputs_of_par trans_of_par)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   319
  apply (tactic "safe_tac (Context.raw_transfer @{theory} @{theory_context Fun})")
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   320
  apply (simp add: inp_is_act)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   321
  prefer 2
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   322
  apply (simp add: inp_is_act)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   323
  text \<open>\<open>a \<in> inp A\<close>\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   324
  apply (case_tac "a \<in> act B")
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   325
  text \<open>\<open>a \<in> inp B\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   326
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   327
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   328
  apply (drule inpAAactB_is_inpBoroutB)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   329
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   330
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   331
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   332
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   333
  apply (erule_tac x = "aa" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   334
  apply (erule_tac x = "b" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   335
  apply (erule exE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   336
  apply (erule exE)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   337
  apply (rule_tac x = "(s2, s2a)" in exI)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   338
  apply (simp add: inp_is_act)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   339
  text \<open>\<open>a \<notin> act B\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   340
  apply (simp add: inp_is_act)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   341
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   342
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   343
  apply (erule_tac x = "aa" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   344
  apply (erule exE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   345
  apply (rule_tac x = " (s2,b) " in exI)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   346
  apply simp
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   347
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   348
  text \<open>\<open>a \<in> inp B\<close>\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   349
  apply (case_tac "a \<in> act A")
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   350
  text \<open>\<open>a \<in> act A\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   351
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   352
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   353
  apply (simp add: inp_is_act)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   354
  apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   355
  apply (drule inpAAactB_is_inpBoroutB)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   356
  back
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   357
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   358
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   359
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   360
  apply (erule_tac x = "aa" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   361
  apply (erule_tac x = "b" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   362
  apply (erule exE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   363
  apply (erule exE)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   364
  apply (rule_tac x = "(s2, s2a)" in exI)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   365
  apply (simp add: inp_is_act)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   366
  text \<open>\<open>a \<notin> act B\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   367
  apply (simp add: inp_is_act)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   368
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   369
  apply (erule_tac x = "a" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   370
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   371
  apply (erule_tac x = "b" in allE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   372
  apply (erule exE)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   373
  apply (rule_tac x = "(aa, s2)" in exI)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   374
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   375
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   376
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   377
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   378
subsection \<open>Invariants\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   379
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   380
lemma invariantI:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   381
  assumes "\<And>s. s \<in> starts_of A \<Longrightarrow> P s"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   382
    and "\<And>s t a. reachable A s \<Longrightarrow> P s \<Longrightarrow> (s, a, t) \<in> trans_of A \<longrightarrow> P t"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   383
  shows "invariant A P"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   384
  using assms
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   385
  apply (unfold invariant_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   386
  apply (rule allI)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   387
  apply (rule impI)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   388
  apply (rule_tac x = "s" in reachable.induct)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   389
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   390
  apply blast
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   391
  apply blast
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   392
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   393
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   394
lemma invariantI1:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   395
  assumes "\<And>s. s \<in> starts_of A \<Longrightarrow> P s"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   396
    and "\<And>s t a. reachable A s \<Longrightarrow> P s \<longrightarrow> (s, a, t) \<in> trans_of A \<longrightarrow> P t"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   397
  shows "invariant A P"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   398
  using assms by (blast intro: invariantI)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   399
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   400
lemma invariantE: "invariant A P \<Longrightarrow> reachable A s \<Longrightarrow> P s"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   401
  unfolding invariant_def by blast
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   402
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   403
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   404
subsection \<open>\<open>restrict\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   405
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   406
lemmas reachable_0 = reachable.reachable_0
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   407
  and reachable_n = reachable.reachable_n
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   408
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   409
lemma cancel_restrict_a:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   410
  "starts_of (restrict ioa acts) = starts_of ioa \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   411
   trans_of (restrict ioa acts) = trans_of ioa"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   412
  by (simp add: restrict_def ioa_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   413
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   414
lemma cancel_restrict_b: "reachable (restrict ioa acts) s = reachable ioa s"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   415
  apply (rule iffI)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   416
  apply (erule reachable.induct)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   417
  apply (simp add: cancel_restrict_a reachable_0)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   418
  apply (erule reachable_n)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   419
  apply (simp add: cancel_restrict_a)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   420
  text \<open>\<open>\<longleftarrow>\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   421
  apply (erule reachable.induct)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   422
  apply (rule reachable_0)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   423
  apply (simp add: cancel_restrict_a)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   424
  apply (erule reachable_n)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   425
  apply (simp add: cancel_restrict_a)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   426
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   427
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   428
lemma acts_restrict: "act (restrict A acts) = act A"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   429
  by (auto simp add: actions_def asig_internals_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   430
    asig_outputs_def asig_inputs_def externals_def asig_of_def restrict_def restrict_asig_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   431
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   432
lemma cancel_restrict:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   433
  "starts_of (restrict ioa acts) = starts_of ioa \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   434
   trans_of (restrict ioa acts) = trans_of ioa \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   435
   reachable (restrict ioa acts) s = reachable ioa s \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   436
   act (restrict A acts) = act A"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   437
  by (simp add: cancel_restrict_a cancel_restrict_b acts_restrict)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   438
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   439
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   440
subsection \<open>\<open>rename\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   441
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   442
lemma trans_rename: "s \<midarrow>a\<midarrow>(rename C f)\<rightarrow> t \<Longrightarrow> (\<exists>x. Some x = f a \<and> s \<midarrow>x\<midarrow>C\<rightarrow> t)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   443
  by (simp add: Let_def rename_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   444
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   445
lemma reachable_rename: "reachable (rename C g) s \<Longrightarrow> reachable C s"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   446
  apply (erule reachable.induct)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   447
  apply (rule reachable_0)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   448
  apply (simp add: rename_def ioa_projections)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   449
  apply (drule trans_rename)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   450
  apply (erule exE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   451
  apply (erule conjE)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   452
  apply (erule reachable_n)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   453
  apply assumption
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   454
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   455
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   456
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   457
subsection \<open>\<open>trans_of (A \<parallel> B)\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   458
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   459
lemma trans_A_proj:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   460
  "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act A \<Longrightarrow> (fst s, a, fst t) \<in> trans_of A"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   461
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   462
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   463
lemma trans_B_proj:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   464
  "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act B \<Longrightarrow> (snd s, a, snd t) \<in> trans_of B"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   465
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   466
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   467
lemma trans_A_proj2: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<notin> act A \<Longrightarrow> fst s = fst t"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   468
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   469
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   470
lemma trans_B_proj2: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<notin> act B \<Longrightarrow> snd s = snd t"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   471
  by (simp add: Let_def par_def trans_of_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   472
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   473
lemma trans_AB_proj: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act A \<or> a \<in> act B"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   474
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   475
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   476
lemma trans_AB:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   477
  "a \<in> act A \<Longrightarrow> a \<in> act B \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   478
  (fst s, a, fst t) \<in> trans_of A \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   479
  (snd s, a, snd t) \<in> trans_of B \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   480
  (s, a, t) \<in> trans_of (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   481
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   482
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   483
lemma trans_A_notB:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   484
  "a \<in> act A \<Longrightarrow> a \<notin> act B \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   485
  (fst s, a, fst t) \<in> trans_of A \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   486
  snd s = snd t \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   487
  (s, a, t) \<in> trans_of (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   488
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   489
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   490
lemma trans_notA_B:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   491
  "a \<notin> act A \<Longrightarrow> a \<in> act B \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   492
  (snd s, a, snd t) \<in> trans_of B \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   493
  fst s = fst t \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   494
  (s, a, t) \<in> trans_of (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   495
  by (simp add: Let_def par_def trans_of_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   496
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   497
lemmas trans_of_defs1 = trans_AB trans_A_notB trans_notA_B
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   498
  and trans_of_defs2 = trans_A_proj trans_B_proj trans_A_proj2 trans_B_proj2 trans_AB_proj
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   499
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   500
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   501
lemma trans_of_par4:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   502
  "(s, a, t) \<in> trans_of (A \<parallel> B \<parallel> C \<parallel> D) \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   503
    ((a \<in> actions (asig_of A) \<or> a \<in> actions (asig_of B) \<or> a \<in> actions (asig_of C) \<or>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   504
      a \<in> actions (asig_of D)) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   505
     (if a \<in> actions (asig_of A)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   506
      then (fst s, a, fst t) \<in> trans_of A
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   507
      else fst t = fst s) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   508
     (if a \<in> actions (asig_of B)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   509
      then (fst (snd s), a, fst (snd t)) \<in> trans_of B
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   510
      else fst (snd t) = fst (snd s)) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   511
     (if a \<in> actions (asig_of C)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   512
      then (fst (snd (snd s)), a, fst (snd (snd t))) \<in> trans_of C
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   513
      else fst (snd (snd t)) = fst (snd (snd s))) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   514
     (if a \<in> actions (asig_of D)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   515
      then (snd (snd (snd s)), a, snd (snd (snd t))) \<in> trans_of D
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   516
      else snd (snd (snd t)) = snd (snd (snd s))))"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   517
  by (simp add: par_def actions_asig_comp prod_eq_iff Let_def ioa_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   518
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   519
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   520
subsection \<open>Proof obligation generator for IOA requirements\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   521
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   522
(*without assumptions on A and B because is_trans_of is also incorporated in par_def*)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   523
lemma is_trans_of_par: "is_trans_of (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   524
  by (simp add: is_trans_of_def Let_def actions_of_par trans_of_par)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   525
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   526
lemma is_trans_of_restrict: "is_trans_of A \<Longrightarrow> is_trans_of (restrict A acts)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   527
  by (simp add: is_trans_of_def cancel_restrict acts_restrict)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   528
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   529
lemma is_trans_of_rename: "is_trans_of A \<Longrightarrow> is_trans_of (rename A f)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   530
  apply (unfold is_trans_of_def restrict_def restrict_asig_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   531
  apply (simp add: Let_def actions_def trans_of_def asig_internals_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   532
    asig_outputs_def asig_inputs_def externals_def asig_of_def rename_def rename_set_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   533
  apply blast
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   534
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   535
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   536
lemma is_asig_of_par: "is_asig_of A \<Longrightarrow> is_asig_of B \<Longrightarrow> compatible A B \<Longrightarrow> is_asig_of (A \<parallel> B)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   537
  apply (simp add: is_asig_of_def asig_of_par asig_comp_def compatible_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   538
    asig_internals_def asig_outputs_def asig_inputs_def actions_def is_asig_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   539
  apply (simp add: asig_of_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   540
  apply auto
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   541
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   542
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   543
lemma is_asig_of_restrict: "is_asig_of A \<Longrightarrow> is_asig_of (restrict A f)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   544
  apply (unfold is_asig_of_def is_asig_def asig_of_def restrict_def restrict_asig_def
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   545
    asig_internals_def asig_outputs_def asig_inputs_def externals_def o_def)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   546
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   547
  apply auto
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   548
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   549
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   550
lemma is_asig_of_rename: "is_asig_of A \<Longrightarrow> is_asig_of (rename A f)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   551
  apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   552
    asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   553
  apply auto
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   554
  apply (drule_tac [!] s = "Some _" in sym)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   555
  apply auto
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   556
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   557
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   558
lemmas [simp] = is_asig_of_par is_asig_of_restrict
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   559
  is_asig_of_rename is_trans_of_par is_trans_of_restrict is_trans_of_rename
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   560
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   561
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   562
lemma compatible_par: "compatible A B \<Longrightarrow> compatible A C \<Longrightarrow> compatible A (B \<parallel> C)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   563
  by (auto simp add: compatible_def internals_of_par outputs_of_par actions_of_par)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   564
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   565
(*better derive by previous one and compat_commute*)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   566
lemma compatible_par2: "compatible A C \<Longrightarrow> compatible B C \<Longrightarrow> compatible (A \<parallel> B) C"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   567
  by (auto simp add: compatible_def internals_of_par outputs_of_par actions_of_par)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   568
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   569
lemma compatible_restrict:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   570
  "compatible A B \<Longrightarrow> (ext B - S) \<inter> ext A = {} \<Longrightarrow> compatible A (restrict B S)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   571
  by (auto simp add: compatible_def ioa_triple_proj asig_triple_proj externals_def
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   572
    restrict_def restrict_asig_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   573
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   574
declare split_paired_Ex [simp]
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   575
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   576
end