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