src/HOL/UNITY/Comp/Alloc.thy
author wenzelm
Mon, 06 Sep 2010 19:13:10 +0200
changeset 39159 0dec18004e75
parent 38012 3ca193a6ae5a
child 39198 f967a16dfcdd
permissions -rw-r--r--
more antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 36866
diff changeset
     1
(*  Title:      HOL/UNITY/Comp/Alloc.thy
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
Specification of Chandy and Charpentier's Allocator
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
*)
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     7
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
     8
theory Alloc
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
     9
imports AllocBase "../PPROD"
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    10
begin
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    11
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    12
subsection{*State definitions.  OUTPUT variables are locals*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    13
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    14
record clientState =
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    15
  giv :: "nat list"   --{*client's INPUT history:  tokens GRANTED*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    16
  ask :: "nat list"   --{*client's OUTPUT history: tokens REQUESTED*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    17
  rel :: "nat list"   --{*client's OUTPUT history: tokens RELEASED*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    18
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    19
record 'a clientState_d =
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    20
  clientState +
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    21
  dummy :: 'a       --{*dummy field for new variables*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    22
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    23
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    24
  --{*DUPLICATED FROM Client.thy, but with "tok" removed*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    25
  --{*Maybe want a special theory section to declare such maps*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    26
  non_dummy :: "'a clientState_d => clientState"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    27
  where "non_dummy s = (|giv = giv s, ask = ask s, rel = rel s|)"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    28
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    29
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    30
  --{*Renaming map to put a Client into the standard form*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    31
  client_map :: "'a clientState_d => clientState*'a"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    32
  where "client_map = funPair non_dummy dummy"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    33
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    34
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    35
record allocState =
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    36
  allocGiv :: "nat => nat list"   --{*OUTPUT history: source of "giv" for i*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    37
  allocAsk :: "nat => nat list"   --{*INPUT: allocator's copy of "ask" for i*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    38
  allocRel :: "nat => nat list"   --{*INPUT: allocator's copy of "rel" for i*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    39
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    40
record 'a allocState_d =
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    41
  allocState +
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    42
  dummy    :: 'a                --{*dummy field for new variables*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    43
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    44
record 'a systemState =
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    45
  allocState +
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    46
  client :: "nat => clientState"  --{*states of all clients*}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    47
  dummy  :: 'a                    --{*dummy field for new variables*}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    48
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    49
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    50
--{** Resource allocation system specification **}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    51
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    52
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    53
  --{*spec (1)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    54
  system_safety :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    55
  where "system_safety =
15074
277b3a4da341 Modified \<Sum> syntax a little.
nipkow
parents: 11786
diff changeset
    56
     Always {s. (SUM i: lessThan Nclients. (tokens o giv o sub i o client)s)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    57
     \<le> NbT + (SUM i: lessThan Nclients. (tokens o rel o sub i o client)s)}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    58
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    59
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    60
  --{*spec (2)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    61
  system_progress :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    62
  where "system_progress = (INT i : lessThan Nclients.
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    63
                        INT h.
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    64
                          {s. h \<le> (ask o sub i o client)s} LeadsTo
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    65
                          {s. h pfixLe (giv o sub i o client) s})"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    66
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    67
definition
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    68
  system_spec :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    69
  where "system_spec = system_safety Int system_progress"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    70
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    71
--{** Client specification (required) ***}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    72
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    73
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    74
  --{*spec (3)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    75
  client_increasing :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    76
  where "client_increasing = UNIV guarantees  Increasing ask Int Increasing rel"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    77
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    78
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    79
  --{*spec (4)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    80
  client_bounded :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    81
  where "client_bounded = UNIV guarantees  Always {s. ALL elt : set (ask s). elt \<le> NbT}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    82
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    83
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    84
  --{*spec (5)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    85
  client_progress :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    86
  where "client_progress =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    87
         Increasing giv  guarantees
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    88
         (INT h. {s. h \<le> giv s & h pfixGe ask s}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
    89
                 LeadsTo {s. tokens h \<le> (tokens o rel) s})"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    90
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    91
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    92
  --{*spec: preserves part*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    93
  client_preserves :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    94
  where "client_preserves = preserves giv Int preserves clientState_d.dummy"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    95
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    96
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
    97
  --{*environmental constraints*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
    98
  client_allowed_acts :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
    99
  where "client_allowed_acts =
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   100
       {F. AllowedActs F =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   101
            insert Id (UNION (preserves (funPair rel ask)) Acts)}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   102
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   103
definition
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   104
  client_spec :: "'a clientState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   105
  where "client_spec = client_increasing Int client_bounded Int client_progress
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   106
                    Int client_allowed_acts Int client_preserves"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   107
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   108
--{** Allocator specification (required) **}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   109
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   110
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   111
  --{*spec (6)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   112
  alloc_increasing :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   113
  where "alloc_increasing =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   114
         UNIV  guarantees
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   115
         (INT i : lessThan Nclients. Increasing (sub i o allocGiv))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   116
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   117
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   118
  --{*spec (7)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   119
  alloc_safety :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   120
  where "alloc_safety =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   121
         (INT i : lessThan Nclients. Increasing (sub i o allocRel))
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   122
         guarantees
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   123
         Always {s. (SUM i: lessThan Nclients. (tokens o sub i o allocGiv)s)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   124
         \<le> NbT + (SUM i: lessThan Nclients. (tokens o sub i o allocRel)s)}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   125
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   126
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   127
  --{*spec (8)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   128
  alloc_progress :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   129
  where "alloc_progress =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   130
         (INT i : lessThan Nclients. Increasing (sub i o allocAsk) Int
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   131
                                     Increasing (sub i o allocRel))
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   132
         Int
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   133
         Always {s. ALL i<Nclients.
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   134
                     ALL elt : set ((sub i o allocAsk) s). elt \<le> NbT}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   135
         Int
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   136
         (INT i : lessThan Nclients.
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   137
          INT h. {s. h \<le> (sub i o allocGiv)s & h pfixGe (sub i o allocAsk)s}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   138
                 LeadsTo
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   139
                 {s. tokens h \<le> (tokens o sub i o allocRel)s})
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   140
         guarantees
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   141
             (INT i : lessThan Nclients.
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   142
              INT h. {s. h \<le> (sub i o allocAsk) s}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   143
                     LeadsTo
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   144
                     {s. h pfixLe (sub i o allocGiv) s})"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   145
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   146
  (*NOTE: to follow the original paper, the formula above should have had
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   147
        INT h. {s. h i \<le> (sub i o allocGiv)s & h i pfixGe (sub i o allocAsk)s}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   148
               LeadsTo
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   149
               {s. tokens h i \<le> (tokens o sub i o allocRel)s})
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   150
    thus h should have been a function variable.  However, only h i is ever
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   151
    looked at.*)
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   152
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   153
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   154
  --{*spec: preserves part*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   155
  alloc_preserves :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   156
  where "alloc_preserves = preserves allocRel Int preserves allocAsk Int
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   157
                        preserves allocState_d.dummy"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   158
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   159
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   160
  --{*environmental constraints*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   161
  alloc_allowed_acts :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   162
  where "alloc_allowed_acts =
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   163
       {F. AllowedActs F =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   164
            insert Id (UNION (preserves allocGiv) Acts)}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   165
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   166
definition
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   167
  alloc_spec :: "'a allocState_d program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   168
  where "alloc_spec = alloc_increasing Int alloc_safety Int alloc_progress Int
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   169
                   alloc_allowed_acts Int alloc_preserves"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   170
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   171
--{** Network specification **}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   172
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   173
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   174
  --{*spec (9.1)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   175
  network_ask :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   176
  where "network_ask = (INT i : lessThan Nclients.
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   177
                        Increasing (ask o sub i o client)  guarantees
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   178
                        ((sub i o allocAsk) Fols (ask o sub i o client)))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   179
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   180
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   181
  --{*spec (9.2)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   182
  network_giv :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   183
  where "network_giv = (INT i : lessThan Nclients.
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   184
                        Increasing (sub i o allocGiv)
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   185
                        guarantees
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   186
                        ((giv o sub i o client) Fols (sub i o allocGiv)))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   187
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   188
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   189
  --{*spec (9.3)*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   190
  network_rel :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   191
  where "network_rel = (INT i : lessThan Nclients.
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   192
                        Increasing (rel o sub i o client)
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   193
                        guarantees
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   194
                        ((sub i o allocRel) Fols (rel o sub i o client)))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   195
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   196
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   197
  --{*spec: preserves part*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   198
  network_preserves :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   199
  where "network_preserves =
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   200
       preserves allocGiv  Int
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   201
       (INT i : lessThan Nclients. preserves (rel o sub i o client)  Int
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   202
                                   preserves (ask o sub i o client))"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   203
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   204
definition
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   205
  --{*environmental constraints*}
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   206
  network_allowed_acts :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   207
  where "network_allowed_acts =
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   208
       {F. AllowedActs F =
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   209
           insert Id
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   210
            (UNION (preserves allocRel Int
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   211
                    (INT i: lessThan Nclients. preserves(giv o sub i o client)))
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   212
                  Acts)}"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   213
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   214
definition
17310
1322ed8e0ee4 converted to Isar theory format;
wenzelm
parents: 15074
diff changeset
   215
  network_spec :: "'a systemState program set"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   216
  where "network_spec = network_ask Int network_giv Int
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   217
                     network_rel Int network_allowed_acts Int
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   218
                     network_preserves"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   219
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   220
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   221
--{** State mappings **}
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   222
definition
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   223
  sysOfAlloc :: "((nat => clientState) * 'a) allocState_d => 'a systemState"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   224
  where "sysOfAlloc = (%s. let (cl,xtr) = allocState_d.dummy s
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   225
                       in (| allocGiv = allocGiv s,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   226
                             allocAsk = allocAsk s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   227
                             allocRel = allocRel s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   228
                             client   = cl,
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   229
                             dummy    = xtr|))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   230
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   231
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   232
definition
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   233
  sysOfClient :: "(nat => clientState) * 'a allocState_d => 'a systemState"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   234
  where "sysOfClient = (%(cl,al). (| allocGiv = allocGiv al,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   235
                                 allocAsk = allocAsk al,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   236
                                 allocRel = allocRel al,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   237
                                 client   = cl,
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   238
                                 systemState.dummy = allocState_d.dummy al|))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   239
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   240
axiomatization Alloc :: "'a allocState_d program"
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   241
  where Alloc: "Alloc : alloc_spec"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   242
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   243
axiomatization Client :: "'a clientState_d program"
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   244
  where Client: "Client : client_spec"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   245
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   246
axiomatization Network :: "'a systemState program"
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   247
  where Network: "Network : network_spec"
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   248
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   249
definition System  :: "'a systemState program"
426d5781bb25 modernized specifications;
wenzelm
parents: 32149
diff changeset
   250
  where "System = rename sysOfAlloc Alloc Join Network Join
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   251
                 (rename sysOfClient
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   252
                  (plam x: lessThan Nclients. rename client_map Client))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   253
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   254
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   255
(**
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   256
locale System =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   257
  fixes
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   258
    Alloc   :: 'a allocState_d program
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   259
    Client  :: 'a clientState_d program
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   260
    Network :: 'a systemState program
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   261
    System  :: 'a systemState program
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   262
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   263
  assumes
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   264
    Alloc   "Alloc   : alloc_spec"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   265
    Client  "Client  : client_spec"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   266
    Network "Network : network_spec"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   267
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   268
  defines
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   269
    System_def
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   270
      "System == rename sysOfAlloc Alloc
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   271
                 Join
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   272
                 Network
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   273
                 Join
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   274
                 (rename sysOfClient
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   275
                  (plam x: lessThan Nclients. rename client_map Client))"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   276
**)
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   277
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   278
(*Perhaps equalities.ML shouldn't add this in the first place!*)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   279
declare image_Collect [simp del]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   280
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   281
declare subset_preserves_o [THEN [2] rev_subsetD, intro]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   282
declare subset_preserves_o [THEN [2] rev_subsetD, simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   283
declare funPair_o_distrib [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   284
declare Always_INT_distrib [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   285
declare o_apply [simp del]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   286
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   287
(*For rewriting of specifications related by "guarantees"*)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   288
lemmas [simp] =
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   289
  rename_image_constrains
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   290
  rename_image_stable
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   291
  rename_image_increasing
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   292
  rename_image_invariant
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   293
  rename_image_Constrains
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   294
  rename_image_Stable
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   295
  rename_image_Increasing
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   296
  rename_image_Always
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   297
  rename_image_leadsTo
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   298
  rename_image_LeadsTo
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   299
  rename_preserves
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   300
  rename_image_preserves
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   301
  lift_image_preserves
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   302
  bij_image_INT
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   303
  bij_is_inj [THEN image_Int]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   304
  bij_image_Collect_eq
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   305
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   306
ML {*
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   307
(*Splits up conjunctions & intersections: like CONJUNCTS in the HOL system*)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   308
fun list_of_Int th =
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   309
    (list_of_Int (th RS conjunct1) @ list_of_Int (th RS conjunct2))
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   310
    handle THM _ => (list_of_Int (th RS IntD1) @ list_of_Int (th RS IntD2))
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   311
    handle THM _ => (list_of_Int (th RS @{thm INT_D}))
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   312
    handle THM _ => (list_of_Int (th RS bspec))
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   313
    handle THM _ => [th];
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   314
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   315
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   316
lemmas lessThanBspec = lessThan_iff [THEN iffD2, THEN [2] bspec]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   317
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   318
attribute_setup normalized = {*
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   319
let
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   320
  fun normalized th =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   321
    normalized (th RS spec
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   322
      handle THM _ => th RS @{thm lessThanBspec}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   323
      handle THM _ => th RS bspec
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   324
      handle THM _ => th RS (@{thm guarantees_INT_right_iff} RS iffD1))
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   325
    handle THM _ => th;
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   326
in
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   327
  Scan.succeed (Thm.rule_attribute (K normalized))
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   328
end
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   329
*} ""
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   330
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   331
(*** bijectivity of sysOfAlloc [MUST BE AUTOMATED] ***)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   332
ML {*
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   333
fun record_auto_tac (cs, ss) =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37936
diff changeset
   334
  auto_tac (cs addIs [ext] addSWrapper Record.split_wrapper,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   335
    ss addsimps [@{thm sysOfAlloc_def}, @{thm sysOfClient_def},
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   336
      @{thm client_map_def}, @{thm non_dummy_def}, @{thm funPair_def},
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   337
      @{thm o_apply}, @{thm Let_def}])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   338
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   339
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   340
method_setup record_auto = {*
32149
ef59550a55d3 renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 32139
diff changeset
   341
  Scan.succeed (fn ctxt => SIMPLE_METHOD (record_auto_tac (clasimpset_of ctxt)))
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   342
*} ""
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   343
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   344
lemma inj_sysOfAlloc [iff]: "inj sysOfAlloc"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   345
  apply (unfold sysOfAlloc_def Let_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   346
  apply (rule inj_onI)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   347
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   348
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   349
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   350
text{*We need the inverse; also having it simplifies the proof of surjectivity*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   351
lemma inv_sysOfAlloc_eq [simp]: "!!s. inv sysOfAlloc s =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   352
             (| allocGiv = allocGiv s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   353
                allocAsk = allocAsk s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   354
                allocRel = allocRel s,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   355
                allocState_d.dummy = (client s, dummy s) |)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   356
  apply (rule inj_sysOfAlloc [THEN inv_f_eq])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   357
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   358
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   359
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   360
lemma surj_sysOfAlloc [iff]: "surj sysOfAlloc"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   361
  apply (simp add: surj_iff expand_fun_eq o_apply)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   362
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   363
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   364
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   365
lemma bij_sysOfAlloc [iff]: "bij sysOfAlloc"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   366
  apply (blast intro: bijI)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   367
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   368
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   369
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   370
subsubsection{*bijectivity of @{term sysOfClient}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   371
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   372
lemma inj_sysOfClient [iff]: "inj sysOfClient"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   373
  apply (unfold sysOfClient_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   374
  apply (rule inj_onI)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   375
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   376
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   377
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   378
lemma inv_sysOfClient_eq [simp]: "!!s. inv sysOfClient s =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   379
             (client s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   380
              (| allocGiv = allocGiv s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   381
                 allocAsk = allocAsk s,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   382
                 allocRel = allocRel s,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   383
                 allocState_d.dummy = systemState.dummy s|) )"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   384
  apply (rule inj_sysOfClient [THEN inv_f_eq])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   385
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   386
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   387
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   388
lemma surj_sysOfClient [iff]: "surj sysOfClient"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   389
  apply (simp add: surj_iff expand_fun_eq o_apply)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   390
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   391
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   392
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   393
lemma bij_sysOfClient [iff]: "bij sysOfClient"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   394
  apply (blast intro: bijI)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   395
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   396
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   397
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   398
subsubsection{*bijectivity of @{term client_map}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   399
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   400
lemma inj_client_map [iff]: "inj client_map"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   401
  apply (unfold inj_on_def)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   402
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   403
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   404
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   405
lemma inv_client_map_eq [simp]: "!!s. inv client_map s =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   406
             (%(x,y).(|giv = giv x, ask = ask x, rel = rel x,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   407
                       clientState_d.dummy = y|)) s"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   408
  apply (rule inj_client_map [THEN inv_f_eq])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   409
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   410
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   411
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   412
lemma surj_client_map [iff]: "surj client_map"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   413
  apply (simp add: surj_iff expand_fun_eq o_apply)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   414
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   415
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   416
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   417
lemma bij_client_map [iff]: "bij client_map"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   418
  apply (blast intro: bijI)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   419
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   420
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   421
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   422
text{*o-simprules for @{term client_map}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   423
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   424
lemma fst_o_client_map: "fst o client_map = non_dummy"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   425
  apply (unfold client_map_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   426
  apply (rule fst_o_funPair)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   427
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   428
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   429
ML {* bind_thms ("fst_o_client_map'", make_o_equivs @{thm fst_o_client_map}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   430
declare fst_o_client_map' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   431
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   432
lemma snd_o_client_map: "snd o client_map = clientState_d.dummy"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   433
  apply (unfold client_map_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   434
  apply (rule snd_o_funPair)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   435
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   436
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   437
ML {* bind_thms ("snd_o_client_map'", make_o_equivs @{thm snd_o_client_map}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   438
declare snd_o_client_map' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   439
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   440
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   441
subsection{*o-simprules for @{term sysOfAlloc} [MUST BE AUTOMATED]*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   442
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   443
lemma client_o_sysOfAlloc: "client o sysOfAlloc = fst o allocState_d.dummy "
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   444
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   445
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   446
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   447
ML {* bind_thms ("client_o_sysOfAlloc'", make_o_equivs @{thm client_o_sysOfAlloc}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   448
declare client_o_sysOfAlloc' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   449
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   450
lemma allocGiv_o_sysOfAlloc_eq: "allocGiv o sysOfAlloc = allocGiv"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   451
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   452
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   453
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   454
ML {* bind_thms ("allocGiv_o_sysOfAlloc_eq'", make_o_equivs @{thm allocGiv_o_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   455
declare allocGiv_o_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   456
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   457
lemma allocAsk_o_sysOfAlloc_eq: "allocAsk o sysOfAlloc = allocAsk"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   458
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   459
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   460
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   461
ML {* bind_thms ("allocAsk_o_sysOfAlloc_eq'", make_o_equivs @{thm allocAsk_o_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   462
declare allocAsk_o_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   463
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   464
lemma allocRel_o_sysOfAlloc_eq: "allocRel o sysOfAlloc = allocRel"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   465
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   466
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   467
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   468
ML {* bind_thms ("allocRel_o_sysOfAlloc_eq'", make_o_equivs @{thm allocRel_o_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   469
declare allocRel_o_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   470
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   471
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   472
subsection{* o-simprules for @{term sysOfClient} [MUST BE AUTOMATED]*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   473
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   474
lemma client_o_sysOfClient: "client o sysOfClient = fst"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   475
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   476
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   477
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   478
ML {* bind_thms ("client_o_sysOfClient'", make_o_equivs @{thm client_o_sysOfClient}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   479
declare client_o_sysOfClient' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   480
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   481
lemma allocGiv_o_sysOfClient_eq: "allocGiv o sysOfClient = allocGiv o snd "
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   482
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   483
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   484
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   485
ML {* bind_thms ("allocGiv_o_sysOfClient_eq'", make_o_equivs @{thm allocGiv_o_sysOfClient_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   486
declare allocGiv_o_sysOfClient_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   487
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   488
lemma allocAsk_o_sysOfClient_eq: "allocAsk o sysOfClient = allocAsk o snd "
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   489
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   490
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   491
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   492
ML {* bind_thms ("allocAsk_o_sysOfClient_eq'", make_o_equivs @{thm allocAsk_o_sysOfClient_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   493
declare allocAsk_o_sysOfClient_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   494
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   495
lemma allocRel_o_sysOfClient_eq: "allocRel o sysOfClient = allocRel o snd "
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   496
  apply record_auto
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   497
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   498
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   499
ML {* bind_thms ("allocRel_o_sysOfClient_eq'", make_o_equivs @{thm allocRel_o_sysOfClient_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   500
declare allocRel_o_sysOfClient_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   501
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   502
lemma allocGiv_o_inv_sysOfAlloc_eq: "allocGiv o inv sysOfAlloc = allocGiv"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   503
  apply (simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   504
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   505
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   506
ML {* bind_thms ("allocGiv_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocGiv_o_inv_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   507
declare allocGiv_o_inv_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   508
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   509
lemma allocAsk_o_inv_sysOfAlloc_eq: "allocAsk o inv sysOfAlloc = allocAsk"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   510
  apply (simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   511
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   512
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   513
ML {* bind_thms ("allocAsk_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocAsk_o_inv_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   514
declare allocAsk_o_inv_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   515
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   516
lemma allocRel_o_inv_sysOfAlloc_eq: "allocRel o inv sysOfAlloc = allocRel"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   517
  apply (simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   518
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   519
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   520
ML {* bind_thms ("allocRel_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocRel_o_inv_sysOfAlloc_eq}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   521
declare allocRel_o_inv_sysOfAlloc_eq' [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   522
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   523
lemma rel_inv_client_map_drop_map: "(rel o inv client_map o drop_map i o inv sysOfClient) =
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   524
      rel o sub i o client"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   525
  apply (simp add: o_def drop_map_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   526
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   527
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   528
ML {* bind_thms ("rel_inv_client_map_drop_map'", make_o_equivs @{thm rel_inv_client_map_drop_map}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   529
declare rel_inv_client_map_drop_map [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   530
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   531
lemma ask_inv_client_map_drop_map: "(ask o inv client_map o drop_map i o inv sysOfClient) =
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   532
      ask o sub i o client"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   533
  apply (simp add: o_def drop_map_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   534
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   535
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   536
ML {* bind_thms ("ask_inv_client_map_drop_map'", make_o_equivs @{thm ask_inv_client_map_drop_map}) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   537
declare ask_inv_client_map_drop_map [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   538
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   539
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   540
declare finite_lessThan [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   541
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   542
text{*Client : <unfolded specification> *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   543
lemmas client_spec_simps =
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   544
  client_spec_def client_increasing_def client_bounded_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   545
  client_progress_def client_allowed_acts_def client_preserves_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   546
  guarantees_Int_right
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   547
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   548
ML {*
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   549
val [Client_Increasing_ask, Client_Increasing_rel,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   550
     Client_Bounded, Client_Progress, Client_AllowedActs,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   551
     Client_preserves_giv, Client_preserves_dummy] =
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   552
        @{thm Client} |> simplify (@{simpset} addsimps @{thms client_spec_simps})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   553
               |> list_of_Int;
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   554
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   555
bind_thm ("Client_Increasing_ask", Client_Increasing_ask);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   556
bind_thm ("Client_Increasing_rel", Client_Increasing_rel);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   557
bind_thm ("Client_Bounded", Client_Bounded);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   558
bind_thm ("Client_Progress", Client_Progress);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   559
bind_thm ("Client_AllowedActs", Client_AllowedActs);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   560
bind_thm ("Client_preserves_giv", Client_preserves_giv);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   561
bind_thm ("Client_preserves_dummy", Client_preserves_dummy);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   562
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   563
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   564
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   565
  Client_Increasing_ask [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   566
  Client_Increasing_rel [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   567
  Client_Bounded [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   568
  Client_preserves_giv [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   569
  Client_preserves_dummy [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   570
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   571
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   572
text{*Network : <unfolded specification> *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   573
lemmas network_spec_simps =
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   574
  network_spec_def network_ask_def network_giv_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   575
  network_rel_def network_allowed_acts_def network_preserves_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   576
  ball_conj_distrib
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   577
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   578
ML {*
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   579
val [Network_Ask, Network_Giv, Network_Rel, Network_AllowedActs,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   580
     Network_preserves_allocGiv, Network_preserves_rel,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   581
     Network_preserves_ask]  =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   582
        @{thm Network} |> simplify (@{simpset} addsimps @{thms network_spec_simps})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   583
                |> list_of_Int;
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   584
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   585
bind_thm ("Network_Ask", Network_Ask);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   586
bind_thm ("Network_Giv", Network_Giv);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   587
bind_thm ("Network_Rel", Network_Rel);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   588
bind_thm ("Network_AllowedActs", Network_AllowedActs);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   589
bind_thm ("Network_preserves_allocGiv", Network_preserves_allocGiv);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   590
bind_thm ("Network_preserves_rel", Network_preserves_rel);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   591
bind_thm ("Network_preserves_ask", Network_preserves_ask);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   592
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   593
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   594
declare Network_preserves_allocGiv [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   595
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   596
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   597
  Network_preserves_rel [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   598
  Network_preserves_ask [simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   599
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   600
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   601
  Network_preserves_rel [simplified o_def, simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   602
  Network_preserves_ask [simplified o_def, simp]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   603
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   604
text{*Alloc : <unfolded specification> *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   605
lemmas alloc_spec_simps =
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   606
  alloc_spec_def alloc_increasing_def alloc_safety_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   607
  alloc_progress_def alloc_allowed_acts_def alloc_preserves_def
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   608
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   609
ML {*
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   610
val [Alloc_Increasing_0, Alloc_Safety, Alloc_Progress, Alloc_AllowedActs,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   611
     Alloc_preserves_allocRel, Alloc_preserves_allocAsk,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   612
     Alloc_preserves_dummy] =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   613
        @{thm Alloc} |> simplify (@{simpset} addsimps @{thms alloc_spec_simps})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   614
              |> list_of_Int;
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   615
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   616
bind_thm ("Alloc_Increasing_0", Alloc_Increasing_0);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   617
bind_thm ("Alloc_Safety", Alloc_Safety);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   618
bind_thm ("Alloc_Progress", Alloc_Progress);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   619
bind_thm ("Alloc_AllowedActs", Alloc_AllowedActs);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   620
bind_thm ("Alloc_preserves_allocRel", Alloc_preserves_allocRel);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   621
bind_thm ("Alloc_preserves_allocAsk", Alloc_preserves_allocAsk);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   622
bind_thm ("Alloc_preserves_dummy", Alloc_preserves_dummy);
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   623
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   624
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   625
text{*Strip off the INT in the guarantees postcondition*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   626
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   627
lemmas Alloc_Increasing = Alloc_Increasing_0 [normalized]
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   628
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   629
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   630
  Alloc_preserves_allocRel [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   631
  Alloc_preserves_allocAsk [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   632
  Alloc_preserves_dummy [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   633
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   634
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   635
subsection{*Components Lemmas [MUST BE AUTOMATED]*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   636
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   637
lemma Network_component_System: "Network Join
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   638
      ((rename sysOfClient
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   639
        (plam x: (lessThan Nclients). rename client_map Client)) Join
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   640
       rename sysOfAlloc Alloc)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   641
      = System"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   642
  by (simp add: System_def Join_ac)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   643
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   644
lemma Client_component_System: "(rename sysOfClient
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   645
       (plam x: (lessThan Nclients). rename client_map Client)) Join
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   646
      (Network Join rename sysOfAlloc Alloc)  =  System"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   647
  by (simp add: System_def Join_ac)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   648
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   649
lemma Alloc_component_System: "rename sysOfAlloc Alloc Join
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   650
       ((rename sysOfClient (plam x: (lessThan Nclients). rename client_map Client)) Join
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   651
        Network)  =  System"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   652
  by (simp add: System_def Join_ac)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   653
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   654
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   655
  Client_component_System [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   656
  Network_component_System [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   657
  Alloc_component_System [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   658
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   659
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   660
text{** These preservation laws should be generated automatically **}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   661
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   662
lemma Client_Allowed [simp]: "Allowed Client = preserves rel Int preserves ask"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   663
  by (auto simp add: Allowed_def Client_AllowedActs safety_prop_Acts_iff)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   664
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   665
lemma Network_Allowed [simp]: "Allowed Network =
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   666
        preserves allocRel Int
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   667
        (INT i: lessThan Nclients. preserves(giv o sub i o client))"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   668
  by (auto simp add: Allowed_def Network_AllowedActs safety_prop_Acts_iff)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   669
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   670
lemma Alloc_Allowed [simp]: "Allowed Alloc = preserves allocGiv"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   671
  by (auto simp add: Allowed_def Alloc_AllowedActs safety_prop_Acts_iff)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   672
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   673
text{*needed in @{text rename_client_map_tac}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   674
lemma OK_lift_rename_Client [simp]: "OK I (%i. lift i (rename client_map Client))"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   675
  apply (rule OK_lift_I)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   676
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   677
  apply (drule_tac w1 = rel in subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   678
  apply (drule_tac [2] w1 = ask in subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   679
  apply (auto simp add: o_def split_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   680
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   681
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   682
lemma fst_lift_map_eq_fst [simp]: "fst (lift_map i x) i = fst x"
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   683
apply (insert fst_o_lift_map [of i])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   684
apply (drule fun_cong [where x=x])
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   685
apply (simp add: o_def);
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   686
done
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   687
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   688
lemma fst_o_lift_map' [simp]:
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   689
     "(f \<circ> sub i \<circ> fst \<circ> lift_map i \<circ> g) = f o fst o g"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   690
apply (subst fst_o_lift_map [symmetric])
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   691
apply (simp only: o_assoc)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   692
done
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   693
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   694
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   695
(*The proofs of rename_Client_Increasing, rename_Client_Bounded and
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   696
  rename_Client_Progress are similar.  All require copying out the original
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   697
  Client property.  A forward proof can be constructed as follows:
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   698
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   699
  Client_Increasing_ask RS
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   700
      (bij_client_map RS rename_rename_guarantees_eq RS iffD2)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   701
  RS (lift_lift_guarantees_eq RS iffD2)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   702
  RS guarantees_PLam_I
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   703
  RS (bij_sysOfClient RS rename_rename_guarantees_eq RS iffD2)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   704
  |> simplify (simpset() addsimps [lift_image_eq_rename, o_def, split_def,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   705
                                   surj_rename RS surj_range])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   706
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   707
However, the "preserves" property remains to be discharged, and the unfolding
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   708
of "o" and "sub" complicates subsequent reasoning.
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   709
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   710
The following tactic works for all three proofs, though it certainly looks
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   711
ad-hoc!
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   712
*)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   713
ML
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   714
{*
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   715
fun rename_client_map_tac ss =
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   716
  EVERY [
25995
21b51f748daf rename_client_map_tac: avoid ill-defined thm reference;
wenzelm
parents: 24147
diff changeset
   717
    simp_tac (ss addsimps [@{thm rename_guarantees_eq_rename_inv}]) 1,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   718
    rtac @{thm guarantees_PLam_I} 1,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   719
    assume_tac 2,
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   720
         (*preserves: routine reasoning*)
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   721
    asm_simp_tac (ss addsimps [@{thm lift_preserves_sub}]) 2,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   722
         (*the guarantee for  "lift i (rename client_map Client)" *)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   723
    asm_simp_tac
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   724
        (ss addsimps [@{thm lift_guarantees_eq_lift_inv},
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   725
                      @{thm rename_guarantees_eq_rename_inv},
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   726
                      @{thm bij_imp_bij_inv}, @{thm surj_rename} RS @{thm surj_range},
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   727
                      @{thm inv_inv_eq}]) 1,
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   728
    asm_simp_tac
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   729
        (@{simpset} addsimps [@{thm o_def}, @{thm non_dummy_def}, @{thm guarantees_Int_right}]) 1]
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   730
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   731
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   732
method_setup rename_client_map = {*
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30528
diff changeset
   733
  Scan.succeed (fn ctxt =>
32149
ef59550a55d3 renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 32139
diff changeset
   734
    SIMPLE_METHOD (rename_client_map_tac (simpset_of ctxt)))
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   735
*} ""
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   736
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   737
text{*Lifting @{text Client_Increasing} to @{term systemState}*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   738
lemma rename_Client_Increasing: "i : I
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   739
      ==> rename sysOfClient (plam x: I. rename client_map Client) :
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   740
            UNIV  guarantees
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   741
            Increasing (ask o sub i o client) Int
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   742
            Increasing (rel o sub i o client)"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   743
  by rename_client_map
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   744
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   745
lemma preserves_sub_fst_lift_map: "[| F : preserves w; i ~= j |]
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   746
      ==> F : preserves (sub i o fst o lift_map j o funPair v w)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   747
  apply (auto simp add: lift_map_def split_def linorder_neq_iff o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   748
  apply (drule_tac [!] subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   749
  apply (auto simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   750
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   751
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   752
lemma client_preserves_giv_oo_client_map: "[| i < Nclients; j < Nclients |]
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   753
      ==> Client : preserves (giv o sub i o fst o lift_map j o client_map)"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   754
  apply (case_tac "i=j")
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   755
  apply (simp, simp add: o_def non_dummy_def)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   756
  apply (drule Client_preserves_dummy [THEN preserves_sub_fst_lift_map])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   757
  apply (drule_tac [!] subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   758
  apply (simp add: o_def client_map_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   759
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   760
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   761
lemma rename_sysOfClient_ok_Network:
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   762
  "rename sysOfClient (plam x: lessThan Nclients. rename client_map Client)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   763
    ok Network"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   764
  by (auto simp add: ok_iff_Allowed client_preserves_giv_oo_client_map)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   765
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   766
lemma rename_sysOfClient_ok_Alloc:
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   767
  "rename sysOfClient (plam x: lessThan Nclients. rename client_map Client)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   768
    ok rename sysOfAlloc Alloc"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   769
  by (simp add: ok_iff_Allowed)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   770
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   771
lemma rename_sysOfAlloc_ok_Network: "rename sysOfAlloc Alloc ok Network"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   772
  by (simp add: ok_iff_Allowed)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   773
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   774
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   775
  rename_sysOfClient_ok_Network [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   776
  rename_sysOfClient_ok_Alloc [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   777
  rename_sysOfAlloc_ok_Network [iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   778
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   779
text{*The "ok" laws, re-oriented.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   780
  But not sure this works: theorem @{text ok_commute} is needed below*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   781
declare
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   782
  rename_sysOfClient_ok_Network [THEN ok_sym, iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   783
  rename_sysOfClient_ok_Alloc [THEN ok_sym, iff]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   784
  rename_sysOfAlloc_ok_Network [THEN ok_sym]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   785
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   786
lemma System_Increasing: "i < Nclients
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   787
      ==> System : Increasing (ask o sub i o client) Int
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   788
                   Increasing (rel o sub i o client)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   789
  apply (rule component_guaranteesD [OF rename_Client_Increasing Client_component_System])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   790
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   791
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   792
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   793
lemmas rename_guarantees_sysOfAlloc_I =
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   794
  bij_sysOfAlloc [THEN rename_rename_guarantees_eq, THEN iffD2, standard]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   795
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   796
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   797
(*Lifting Alloc_Increasing up to the level of systemState*)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   798
lemmas rename_Alloc_Increasing =
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   799
  Alloc_Increasing
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   800
    [THEN rename_guarantees_sysOfAlloc_I,
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   801
     simplified surj_rename [THEN surj_range] o_def sub_apply
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   802
                rename_image_Increasing bij_sysOfAlloc
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   803
                allocGiv_o_inv_sysOfAlloc_eq'];
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   804
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   805
lemma System_Increasing_allocGiv:
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   806
     "i < Nclients ==> System : Increasing (sub i o allocGiv)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   807
  apply (unfold System_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   808
  apply (simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   809
  apply (rule rename_Alloc_Increasing [THEN guarantees_Join_I1, THEN guaranteesD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   810
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   811
  done
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   812
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   813
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   814
ML {*
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   815
bind_thms ("System_Increasing'", list_of_Int @{thm System_Increasing})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   816
*}
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   817
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   818
declare System_Increasing' [intro!]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   819
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   820
text{* Follows consequences.
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   821
    The "Always (INT ...) formulation expresses the general safety property
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   822
    and allows it to be combined using @{text Always_Int_rule} below. *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   823
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   824
lemma System_Follows_rel:
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   825
  "i < Nclients ==> System : ((sub i o allocRel) Fols (rel o sub i o client))"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   826
  apply (auto intro!: Network_Rel [THEN component_guaranteesD])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   827
  apply (simp add: ok_commute [of Network])
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   828
  done
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   829
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   830
lemma System_Follows_ask:
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   831
  "i < Nclients ==> System : ((sub i o allocAsk) Fols (ask o sub i o client))"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   832
  apply (auto intro!: Network_Ask [THEN component_guaranteesD])
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   833
  apply (simp add: ok_commute [of Network])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   834
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   835
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   836
lemma System_Follows_allocGiv:
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   837
  "i < Nclients ==> System : (giv o sub i o client) Fols (sub i o allocGiv)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   838
  apply (auto intro!: Network_Giv [THEN component_guaranteesD]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   839
    rename_Alloc_Increasing [THEN component_guaranteesD])
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   840
  apply (simp_all add: o_def non_dummy_def ok_commute [of Network])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   841
  apply (auto intro!: rename_Alloc_Increasing [THEN component_guaranteesD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   842
  done
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   843
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   844
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   845
lemma Always_giv_le_allocGiv: "System : Always (INT i: lessThan Nclients.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   846
                       {s. (giv o sub i o client) s \<le> (sub i o allocGiv) s})"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   847
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   848
  apply (erule System_Follows_allocGiv [THEN Follows_Bounded])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   849
  done
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   850
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   851
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   852
lemma Always_allocAsk_le_ask: "System : Always (INT i: lessThan Nclients.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   853
                       {s. (sub i o allocAsk) s \<le> (ask o sub i o client) s})"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   854
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   855
  apply (erule System_Follows_ask [THEN Follows_Bounded])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   856
  done
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   857
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   858
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   859
lemma Always_allocRel_le_rel: "System : Always (INT i: lessThan Nclients.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   860
                       {s. (sub i o allocRel) s \<le> (rel o sub i o client) s})"
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   861
  by (auto intro!: Follows_Bounded System_Follows_rel)
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   862
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   863
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   864
subsection{*Proof of the safety property (1)*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   865
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   866
text{*safety (1), step 1 is @{text System_Follows_rel}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   867
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   868
text{*safety (1), step 2*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   869
(* i < Nclients ==> System : Increasing (sub i o allocRel) *)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   870
lemmas System_Increasing_allocRel = System_Follows_rel [THEN Follows_Increasing1, standard]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   871
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   872
(*Lifting Alloc_safety up to the level of systemState.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   873
  Simplifying with o_def gets rid of the translations but it unfortunately
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   874
  gets rid of the other "o"s too.*)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   875
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   876
text{*safety (1), step 3*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   877
lemma System_sum_bounded:
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   878
    "System : Always {s. (\<Sum>i \<in> lessThan Nclients. (tokens o sub i o allocGiv) s)
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   879
            \<le> NbT + (\<Sum>i \<in> lessThan Nclients. (tokens o sub i o allocRel) s)}"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   880
  apply (simp add: o_apply)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   881
  apply (insert Alloc_Safety [THEN rename_guarantees_sysOfAlloc_I])
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   882
  apply (simp add: o_def);
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   883
  apply (erule component_guaranteesD)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   884
  apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   885
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   886
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   887
text{* Follows reasoning*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   888
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   889
lemma Always_tokens_giv_le_allocGiv: "System : Always (INT i: lessThan Nclients.
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   890
                          {s. (tokens o giv o sub i o client) s
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   891
                           \<le> (tokens o sub i o allocGiv) s})"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   892
  apply (rule Always_giv_le_allocGiv [THEN Always_weaken])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   893
  apply (auto intro: tokens_mono_prefix simp add: o_apply)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   894
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   895
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   896
lemma Always_tokens_allocRel_le_rel: "System : Always (INT i: lessThan Nclients.
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   897
                          {s. (tokens o sub i o allocRel) s
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   898
                           \<le> (tokens o rel o sub i o client) s})"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   899
  apply (rule Always_allocRel_le_rel [THEN Always_weaken])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   900
  apply (auto intro: tokens_mono_prefix simp add: o_apply)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   901
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   902
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   903
text{*safety (1), step 4 (final result!) *}
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   904
theorem System_safety: "System : system_safety"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   905
  apply (unfold system_safety_def)
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   906
  apply (tactic {* rtac (Always_Int_rule [@{thm System_sum_bounded},
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   907
    @{thm Always_tokens_giv_le_allocGiv}, @{thm Always_tokens_allocRel_le_rel}] RS
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   908
    @{thm Always_weaken}) 1 *})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   909
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   910
  apply (rule setsum_fun_mono [THEN order_trans])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   911
  apply (drule_tac [2] order_trans)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   912
  apply (rule_tac [2] add_le_mono [OF order_refl setsum_fun_mono])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   913
  prefer 3 apply assumption
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   914
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   915
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   916
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   917
subsection {* Proof of the progress property (2) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   918
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   919
text{*progress (2), step 1 is @{text System_Follows_ask} and
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   920
      @{text System_Follows_rel}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   921
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   922
text{*progress (2), step 2; see also @{text System_Increasing_allocRel}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   923
(* i < Nclients ==> System : Increasing (sub i o allocAsk) *)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   924
lemmas System_Increasing_allocAsk =  System_Follows_ask [THEN Follows_Increasing1, standard]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   925
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   926
text{*progress (2), step 3: lifting @{text Client_Bounded} to systemState*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   927
lemma rename_Client_Bounded: "i : I
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   928
    ==> rename sysOfClient (plam x: I. rename client_map Client) :
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   929
          UNIV  guarantees
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   930
          Always {s. ALL elt : set ((ask o sub i o client) s). elt \<le> NbT}"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   931
  by rename_client_map
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   932
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   933
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   934
lemma System_Bounded_ask: "i < Nclients
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   935
      ==> System : Always
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   936
                    {s. ALL elt : set ((ask o sub i o client) s). elt \<le> NbT}"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   937
  apply (rule component_guaranteesD [OF rename_Client_Bounded Client_component_System])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   938
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   939
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   940
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   941
lemma Collect_all_imp_eq: "{x. ALL y. P y --> Q x y} = (INT y: {y. P y}. {x. Q x y})"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   942
  apply blast
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   943
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   944
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   945
text{*progress (2), step 4*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   946
lemma System_Bounded_allocAsk: "System : Always {s. ALL i<Nclients.
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   947
                          ALL elt : set ((sub i o allocAsk) s). elt \<le> NbT}"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   948
  apply (auto simp add: Collect_all_imp_eq)
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   949
  apply (tactic {* rtac (Always_Int_rule [@{thm Always_allocAsk_le_ask},
0dec18004e75 more antiquotations;
wenzelm
parents: 38012
diff changeset
   950
    @{thm System_Bounded_ask}] RS @{thm Always_weaken}) 1 *})
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   951
  apply (auto dest: set_mono)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   952
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   953
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   954
text{*progress (2), step 5 is @{text System_Increasing_allocGiv}*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   955
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   956
text{*progress (2), step 6*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   957
(* i < Nclients ==> System : Increasing (giv o sub i o client) *)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   958
lemmas System_Increasing_giv =  System_Follows_allocGiv [THEN Follows_Increasing1, standard]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   959
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   960
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   961
lemma rename_Client_Progress: "i: I
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   962
   ==> rename sysOfClient (plam x: I. rename client_map Client)
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   963
        : Increasing (giv o sub i o client)
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   964
          guarantees
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   965
          (INT h. {s. h \<le> (giv o sub i o client) s &
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   966
                            h pfixGe (ask o sub i o client) s}
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   967
                  LeadsTo {s. tokens h \<le> (tokens o rel o sub i o client) s})"
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   968
  apply rename_client_map
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   969
  apply (simp add: Client_Progress [simplified o_def])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   970
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   971
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   972
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   973
text{*progress (2), step 7*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   974
lemma System_Client_Progress:
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   975
  "System : (INT i : (lessThan Nclients).
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   976
            INT h. {s. h \<le> (giv o sub i o client) s &
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   977
                       h pfixGe (ask o sub i o client) s}
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   978
                LeadsTo {s. tokens h \<le> (tokens o rel o sub i o client) s})"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   979
  apply (rule INT_I)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   980
(*Couldn't have just used Auto_tac since the "INT h" must be kept*)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   981
  apply (rule component_guaranteesD [OF rename_Client_Progress Client_component_System])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   982
  apply (auto simp add: System_Increasing_giv)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   983
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   984
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   985
(*Concludes
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   986
 System : {s. k \<le> (sub i o allocGiv) s}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   987
          LeadsTo
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   988
          {s. (sub i o allocAsk) s \<le> (ask o sub i o client) s} Int
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   989
          {s. k \<le> (giv o sub i o client) s} *)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   990
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   991
lemmas System_lemma1 =
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   992
  Always_LeadsToD [OF System_Follows_ask [THEN Follows_Bounded]
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   993
                      System_Follows_allocGiv [THEN Follows_LeadsTo]]
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
   994
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   995
lemmas System_lemma2 =
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   996
  PSP_Stable [OF System_lemma1
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
   997
              System_Follows_ask [THEN Follows_Increasing1, THEN IncreasingD]]
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   998
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
   999
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1000
lemma System_lemma3: "i < Nclients
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1001
      ==> System : {s. h \<le> (sub i o allocGiv) s &
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1002
                       h pfixGe (sub i o allocAsk) s}
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1003
                   LeadsTo
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1004
                   {s. h \<le> (giv o sub i o client) s &
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1005
                       h pfixGe (ask o sub i o client) s}"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1006
  apply (rule single_LeadsTo_I)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1007
  apply (rule_tac k6 = "h" and x2 = " (sub i o allocAsk) s"
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1008
         in System_lemma2 [THEN LeadsTo_weaken])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1009
  apply auto
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1010
  apply (blast intro: trans_Ge [THEN trans_genPrefix, THEN transD] prefix_imp_pfixGe)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1011
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1012
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1013
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1014
text{*progress (2), step 8: Client i's "release" action is visible system-wide*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1015
lemma System_Alloc_Client_Progress: "i < Nclients
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1016
      ==> System : {s. h \<le> (sub i o allocGiv) s &
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1017
                       h pfixGe (sub i o allocAsk) s}
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1018
                   LeadsTo {s. tokens h \<le> (tokens o sub i o allocRel) s}"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1019
  apply (rule LeadsTo_Trans)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1020
   prefer 2
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1021
   apply (drule System_Follows_rel [THEN
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1022
     mono_tokens [THEN mono_Follows_o, THEN [2] rev_subsetD],
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1023
     THEN Follows_LeadsTo])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1024
   apply (simp add: o_assoc)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1025
  apply (rule LeadsTo_Trans)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1026
   apply (cut_tac [2] System_Client_Progress)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1027
   prefer 2
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1028
   apply (blast intro: LeadsTo_Basis)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1029
  apply (erule System_lemma3)
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1030
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1031
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1032
text{*Lifting @{text Alloc_Progress} up to the level of systemState*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1033
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1034
text{*progress (2), step 9*}
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1035
lemma System_Alloc_Progress:
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1036
 "System : (INT i : (lessThan Nclients).
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1037
            INT h. {s. h \<le> (sub i o allocAsk) s}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1038
                   LeadsTo {s. h pfixLe (sub i o allocGiv) s})"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1039
  apply (simp only: o_apply sub_def)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1040
  apply (insert Alloc_Progress [THEN rename_guarantees_sysOfAlloc_I])
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 30549
diff changeset
  1041
  apply (simp add: o_def del: INT_iff)
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1042
  apply (erule component_guaranteesD)
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1043
  apply (auto simp add:
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1044
    System_Increasing_allocRel [simplified sub_apply o_def]
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1045
    System_Increasing_allocAsk [simplified sub_apply o_def]
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1046
    System_Bounded_allocAsk [simplified sub_apply o_def]
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1047
    System_Alloc_Client_Progress [simplified sub_apply o_def])
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1048
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1049
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1050
text{*progress (2), step 10 (final result!) *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1051
lemma System_Progress: "System : system_progress"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1052
  apply (unfold system_progress_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1053
  apply (cut_tac System_Alloc_Progress)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1054
  apply (blast intro: LeadsTo_Trans
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1055
    System_Follows_allocGiv [THEN Follows_LeadsTo_pfixLe]
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1056
    System_Follows_ask [THEN Follows_LeadsTo])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1057
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1058
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1059
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1060
theorem System_correct: "System : system_spec"
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1061
  apply (unfold system_spec_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1062
  apply (blast intro: System_safety System_Progress)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1063
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1064
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1065
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1066
text{* Some obsolete lemmas *}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1067
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1068
lemma non_dummy_eq_o_funPair: "non_dummy = (% (g,a,r). (| giv = g, ask = a, rel = r |)) o
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1069
                              (funPair giv (funPair ask rel))"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1070
  apply (rule ext)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1071
  apply (auto simp add: o_def non_dummy_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1072
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1073
24147
edc90be09ac1 misc cleanup of ML bindings (for multihreading);
wenzelm
parents: 21710
diff changeset
  1074
lemma preserves_non_dummy_eq: "(preserves non_dummy) =
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1075
      (preserves rel Int preserves ask Int preserves giv)"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1076
  apply (simp add: non_dummy_eq_o_funPair)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1077
  apply auto
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1078
    apply (drule_tac w1 = rel in subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1079
    apply (drule_tac [2] w1 = ask in subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1080
    apply (drule_tac [3] w1 = giv in subset_preserves_o [THEN [2] rev_subsetD])
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1081
    apply (auto simp add: o_def)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1082
  done
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1083
21710
4e4b7c801142 patched up the proofs agsin
paulson
parents: 21669
diff changeset
  1084
text{*Could go to Extend.ML*}
21632
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1085
lemma bij_fst_inv_inv_eq: "bij f ==> fst (inv (%(x, u). inv f x) z) = f z"
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1086
  apply (rule fst_inv_equalityI)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1087
   apply (rule_tac f = "%z. (f z, ?h z) " in surjI)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1088
   apply (simp add: bij_is_inj inv_f_f)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1089
  apply (simp add: bij_is_surj surj_f_inv_f)
e7c1f1a77d18 converted legacy ML script;
wenzelm
parents: 17310
diff changeset
  1090
  done
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
  1091
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
  1092
end