author | wenzelm |
Wed, 02 Aug 2000 19:40:14 +0200 | |
changeset 9502 | 50ec59aff389 |
parent 9403 | aad13b59b8d9 |
child 10064 | 1a77667b21ef |
permissions | -rw-r--r-- |
9027 | 1 |
(* Title: HOL/UNITY/AllocImpl |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Implementation of a multiple-client allocator from a single-client allocator |
|
7 |
*) |
|
8 |
||
9 |
AllocImpl = AllocBase + Follows + PPROD + |
|
10 |
||
11 |
||
12 |
(** State definitions. OUTPUT variables are locals **) |
|
13 |
||
14 |
(*Type variable 'b is the type of items being merged*) |
|
15 |
record 'b merge = |
|
16 |
In :: nat => 'b list (*merge's INPUT histories: streams to merge*) |
|
17 |
Out :: 'b list (*merge's OUTPUT history: merged items*) |
|
18 |
iOut :: nat list (*merge's OUTPUT history: origins of merged items*) |
|
19 |
||
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
20 |
record ('a,'b) merge_d = |
9027 | 21 |
'b merge + |
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
22 |
dummy :: 'a (*dummy field for new variables*) |
9027 | 23 |
|
24 |
constdefs |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
25 |
non_dummy :: ('a,'b) merge_d => 'b merge |
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
26 |
"non_dummy s == (|In = In s, Out = Out s, iOut = iOut s|)" |
9027 | 27 |
|
28 |
record 'b distr = |
|
29 |
In :: 'b list (*items to distribute*) |
|
30 |
iIn :: nat list (*destinations of items to distribute*) |
|
31 |
Out :: nat => 'b list (*distributed items*) |
|
32 |
||
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
33 |
record ('a,'b) distr_d = |
9027 | 34 |
'b distr + |
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
35 |
dummy :: 'a (*dummy field for new variables*) |
9027 | 36 |
|
37 |
record allocState = |
|
38 |
giv :: nat list (*OUTPUT history: source of tokens*) |
|
39 |
ask :: nat list (*INPUT: tokens requested from allocator*) |
|
40 |
rel :: nat list (*INPUT: tokens released to allocator*) |
|
41 |
||
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
42 |
record 'a allocState_d = |
9027 | 43 |
allocState + |
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
44 |
dummy :: 'a (*dummy field for new variables*) |
9027 | 45 |
|
46 |
record 'a systemState = |
|
47 |
allocState + |
|
48 |
mergeRel :: nat merge |
|
49 |
mergeAsk :: nat merge |
|
50 |
distr :: nat distr |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
51 |
dummy :: 'a (*dummy field for new variables*) |
9027 | 52 |
|
53 |
||
54 |
constdefs |
|
55 |
||
9105
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
56 |
(** Merge specification (the number of inputs is Nclients) ***) |
9027 | 57 |
|
58 |
(*spec (10)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
59 |
merge_increasing :: ('a,'b) merge_d program set |
9027 | 60 |
"merge_increasing == |
61 |
UNIV guarantees[funPair merge.Out merge.iOut] |
|
62 |
(Increasing merge.Out) Int (Increasing merge.iOut)" |
|
63 |
||
64 |
(*spec (11)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
65 |
merge_eqOut :: ('a,'b) merge_d program set |
9027 | 66 |
"merge_eqOut == |
67 |
UNIV guarantees[funPair merge.Out merge.iOut] |
|
68 |
Always {s. length (merge.Out s) = length (merge.iOut s)}" |
|
69 |
||
70 |
(*spec (12)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
71 |
merge_bounded :: ('a,'b) merge_d program set |
9027 | 72 |
"merge_bounded == |
73 |
UNIV guarantees[merge.iOut] |
|
9105
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
74 |
Always {s. ALL elt : set (merge.iOut s). elt < Nclients}" |
9027 | 75 |
|
76 |
(*spec (13)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
77 |
merge_follows :: ('a,'b) merge_d program set |
9027 | 78 |
"merge_follows == |
79 |
(INT i : lessThan Nclients. Increasing (sub i o merge.In)) |
|
80 |
guarantees[funPair merge.Out merge.iOut] |
|
81 |
(INT i : lessThan Nclients. |
|
82 |
(%s. sublist (merge.Out s) |
|
9105
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
83 |
{k. k < size(merge.iOut s) & merge.iOut s! k = i}) |
9027 | 84 |
Fols (sub i o merge.In))" |
85 |
||
86 |
(*spec: preserves part*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
87 |
merge_preserves :: ('a,'b) merge_d program set |
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
88 |
"merge_preserves == preserves (funPair merge.In merge_d.dummy)" |
9027 | 89 |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
90 |
merge_spec :: ('a,'b) merge_d program set |
9027 | 91 |
"merge_spec == merge_increasing Int merge_eqOut Int merge_bounded Int |
92 |
merge_follows Int merge_preserves" |
|
93 |
||
9105
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
94 |
(** Distributor specification (the number of outputs is Nclients) ***) |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
95 |
|
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
96 |
(*spec (14)*) |
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
97 |
distr_follows :: ('a,'b) distr_d program set |
9105
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
98 |
"distr_follows == |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
99 |
Increasing distr.In Int Increasing distr.iIn Int |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
100 |
Always {s. ALL elt : set (distr.iIn s). elt < Nclients} |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
101 |
guarantees[distr.Out] |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
102 |
(INT i : lessThan Nclients. |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
103 |
(sub i o distr.Out) Fols |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
104 |
(%s. sublist (distr.In s) |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
105 |
{k. k < size(distr.iIn s) & distr.iIn s ! k = i}))" |
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
106 |
|
d9257992b845
fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents:
9027
diff
changeset
|
107 |
|
9027 | 108 |
(** Single-client allocator specification (required) ***) |
109 |
||
110 |
(*spec (18)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
111 |
alloc_increasing :: 'a allocState_d program set |
9027 | 112 |
"alloc_increasing == UNIV guarantees[giv] Increasing giv" |
113 |
||
114 |
(*spec (19)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
115 |
alloc_safety :: 'a allocState_d program set |
9027 | 116 |
"alloc_safety == |
117 |
Increasing rel |
|
118 |
guarantees[giv] Always {s. tokens (giv s) <= NbT + tokens (rel s)}" |
|
119 |
||
120 |
(*spec (20)*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
121 |
alloc_progress :: 'a allocState_d program set |
9027 | 122 |
"alloc_progress == |
123 |
Increasing ask Int Increasing rel Int |
|
124 |
Always {s. ALL elt : set (ask s). elt <= NbT} |
|
125 |
Int |
|
126 |
(INT h. {s. h <= giv s & h pfixGe (ask s)} |
|
127 |
LeadsTo |
|
128 |
{s. tokens h <= tokens (rel s)}) |
|
129 |
guarantees[giv] |
|
130 |
(INT h. {s. h <= ask s} LeadsTo {s. h pfixLe giv s})" |
|
131 |
||
132 |
(*spec: preserves part*) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
133 |
alloc_preserves :: 'a allocState_d program set |
9027 | 134 |
"alloc_preserves == preserves (funPair rel |
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
135 |
(funPair ask allocState_d.dummy))" |
9027 | 136 |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
137 |
alloc_spec :: 'a allocState_d program set |
9027 | 138 |
"alloc_spec == alloc_increasing Int alloc_safety Int alloc_progress Int |
139 |
alloc_preserves" |
|
140 |
||
141 |
(**** |
|
142 |
(** Network specification ***) |
|
143 |
||
144 |
(*spec (9.1)*) |
|
145 |
network_ask :: 'a systemState program set |
|
146 |
"network_ask == INT i : lessThan Nclients. |
|
147 |
Increasing (ask o sub i o client) |
|
148 |
guarantees[ask] |
|
149 |
(ask Fols (ask o sub i o client))" |
|
150 |
||
151 |
(*spec (9.2)*) |
|
152 |
network_giv :: 'a systemState program set |
|
153 |
"network_giv == INT i : lessThan Nclients. |
|
154 |
Increasing giv |
|
155 |
guarantees[giv o sub i o client] |
|
156 |
((giv o sub i o client) Fols giv )" |
|
157 |
||
158 |
(*spec (9.3)*) |
|
159 |
network_rel :: 'a systemState program set |
|
160 |
"network_rel == INT i : lessThan Nclients. |
|
161 |
Increasing (rel o sub i o client) |
|
162 |
guarantees[rel] |
|
163 |
(rel Fols (rel o sub i o client))" |
|
164 |
||
165 |
(*spec: preserves part*) |
|
166 |
network_preserves :: 'a systemState program set |
|
167 |
"network_preserves == preserves giv Int |
|
168 |
(INT i : lessThan Nclients. |
|
169 |
preserves (funPair rel ask o sub i o client))" |
|
170 |
||
171 |
network_spec :: 'a systemState program set |
|
172 |
"network_spec == network_ask Int network_giv Int |
|
173 |
network_rel Int network_preserves" |
|
174 |
||
175 |
||
176 |
(** State mappings **) |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
177 |
sysOfAlloc :: "((nat => merge) * 'a) allocState_d => 'a systemState" |
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
178 |
"sysOfAlloc == %s. let (cl,xtr) = allocState_d.dummy s |
9027 | 179 |
in (| giv = giv s, |
180 |
ask = ask s, |
|
181 |
rel = rel s, |
|
182 |
client = cl, |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
183 |
dummy = xtr|)" |
9027 | 184 |
|
185 |
||
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
186 |
sysOfClient :: "(nat => merge) * 'a allocState_d => 'a systemState" |
9027 | 187 |
"sysOfClient == %(cl,al). (| giv = giv al, |
188 |
ask = ask al, |
|
189 |
rel = rel al, |
|
190 |
client = cl, |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
191 |
systemState.dummy = allocState_d.dummy al|)" |
9027 | 192 |
****) |
193 |
||
194 |
consts |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
195 |
Alloc :: 'a allocState_d program |
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
196 |
Merge :: ('a,'b) merge_d program |
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
197 |
|
9027 | 198 |
(* |
199 |
Network :: 'a systemState program |
|
200 |
System :: 'a systemState program |
|
201 |
*) |
|
202 |
||
203 |
rules |
|
204 |
Alloc "Alloc : alloc_spec" |
|
205 |
Merge "Merge : merge_spec" |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
9105
diff
changeset
|
206 |
|
9027 | 207 |
(** Network "Network : network_spec"**) |
208 |
||
209 |
||
210 |
||
211 |
end |