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