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