author | wenzelm |
Sun, 12 Aug 2018 14:28:28 +0200 | |
changeset 68743 | 91162dd89571 |
parent 67717 | 5a1b299fe4af |
child 69313 | b021008c5397 |
permissions | -rw-r--r-- |
37936 | 1 |
(* Title: HOL/UNITY/Comp/Alloc.thy |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
3 |
Copyright 1998 University of Cambridge |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
4 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
5 |
Specification of Chandy and Charpentier's Allocator |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
6 |
*) |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
7 |
|
17310 | 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 |
|
63146 | 12 |
subsection\<open>State definitions. OUTPUT variables are locals\<close> |
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 = |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
15 |
giv :: "nat list" \<comment> \<open>client's INPUT history: tokens GRANTED\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
16 |
ask :: "nat list" \<comment> \<open>client's OUTPUT history: tokens REQUESTED\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
17 |
rel :: "nat list" \<comment> \<open>client's OUTPUT history: tokens RELEASED\<close> |
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 + |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
21 |
dummy :: 'a \<comment> \<open>dummy field for new variables\<close> |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
22 |
|
36866 | 23 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
24 |
\<comment> \<open>DUPLICATED FROM Client.thy, but with "tok" removed\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
25 |
\<comment> \<open>Maybe want a special theory section to declare such maps\<close> |
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 |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
30 |
\<comment> \<open>Renaming map to put a Client into the standard form\<close> |
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 = |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
36 |
allocGiv :: "nat => nat list" \<comment> \<open>OUTPUT history: source of "giv" for i\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
37 |
allocAsk :: "nat => nat list" \<comment> \<open>INPUT: allocator's copy of "ask" for i\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
38 |
allocRel :: "nat => nat list" \<comment> \<open>INPUT: allocator's copy of "rel" for i\<close> |
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 + |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
42 |
dummy :: 'a \<comment> \<open>dummy field for new variables\<close> |
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 + |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
46 |
client :: "nat => clientState" \<comment> \<open>states of all clients\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
47 |
dummy :: 'a \<comment> \<open>dummy field for new variables\<close> |
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 |
|
67445 | 50 |
subsubsection \<open>Resource allocation system specification\<close> |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
51 |
|
36866 | 52 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
53 |
\<comment> \<open>spec (1)\<close> |
17310 | 54 |
system_safety :: "'a systemState program set" |
36866 | 55 |
where "system_safety = |
61954 | 56 |
Always {s. (\<Sum>i \<in> lessThan Nclients. (tokens o giv o sub i o client)s) |
57 |
\<le> NbT + (\<Sum>i \<in> 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 |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
60 |
\<comment> \<open>spec (2)\<close> |
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 |
|
67445 | 71 |
|
72 |
subsubsection \<open>Client specification (required)\<close> |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
73 |
|
36866 | 74 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
75 |
\<comment> \<open>spec (3)\<close> |
17310 | 76 |
client_increasing :: "'a clientState_d program set" |
36866 | 77 |
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
|
78 |
|
36866 | 79 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
80 |
\<comment> \<open>spec (4)\<close> |
17310 | 81 |
client_bounded :: "'a clientState_d program set" |
67613 | 82 |
where "client_bounded = UNIV guarantees Always {s. \<forall>elt \<in> set (ask s). elt \<le> NbT}" |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
83 |
|
36866 | 84 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
85 |
\<comment> \<open>spec (5)\<close> |
17310 | 86 |
client_progress :: "'a clientState_d program set" |
36866 | 87 |
where "client_progress = |
24147 | 88 |
Increasing giv guarantees |
89 |
(INT h. {s. h \<le> giv s & h pfixGe ask s} |
|
90 |
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
|
91 |
|
36866 | 92 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
93 |
\<comment> \<open>spec: preserves part\<close> |
17310 | 94 |
client_preserves :: "'a clientState_d program set" |
36866 | 95 |
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
|
96 |
|
36866 | 97 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
98 |
\<comment> \<open>environmental constraints\<close> |
17310 | 99 |
client_allowed_acts :: "'a clientState_d program set" |
36866 | 100 |
where "client_allowed_acts = |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
101 |
{F. AllowedActs F = |
24147 | 102 |
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
|
103 |
|
36866 | 104 |
definition |
17310 | 105 |
client_spec :: "'a clientState_d program set" |
36866 | 106 |
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
|
107 |
Int client_allowed_acts Int client_preserves" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
108 |
|
67445 | 109 |
|
110 |
subsubsection \<open>Allocator specification (required)\<close> |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
111 |
|
36866 | 112 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
113 |
\<comment> \<open>spec (6)\<close> |
17310 | 114 |
alloc_increasing :: "'a allocState_d program set" |
36866 | 115 |
where "alloc_increasing = |
24147 | 116 |
UNIV guarantees |
117 |
(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
|
118 |
|
36866 | 119 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
120 |
\<comment> \<open>spec (7)\<close> |
17310 | 121 |
alloc_safety :: "'a allocState_d program set" |
36866 | 122 |
where "alloc_safety = |
24147 | 123 |
(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
|
124 |
guarantees |
61954 | 125 |
Always {s. (\<Sum>i \<in> lessThan Nclients. (tokens o sub i o allocGiv)s) |
126 |
\<le> NbT + (\<Sum>i \<in> 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
|
127 |
|
36866 | 128 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
129 |
\<comment> \<open>spec (8)\<close> |
17310 | 130 |
alloc_progress :: "'a allocState_d program set" |
36866 | 131 |
where "alloc_progress = |
24147 | 132 |
(INT i : lessThan Nclients. Increasing (sub i o allocAsk) Int |
133 |
Increasing (sub i o allocRel)) |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
134 |
Int |
67613 | 135 |
Always {s. \<forall>i<Nclients. |
136 |
\<forall>elt \<in> 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
|
137 |
Int |
24147 | 138 |
(INT i : lessThan Nclients. |
139 |
INT h. {s. h \<le> (sub i o allocGiv)s & h pfixGe (sub i o allocAsk)s} |
|
140 |
LeadsTo |
|
141 |
{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
|
142 |
guarantees |
24147 | 143 |
(INT i : lessThan Nclients. |
144 |
INT h. {s. h \<le> (sub i o allocAsk) s} |
|
145 |
LeadsTo |
|
146 |
{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
|
147 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
148 |
(*NOTE: to follow the original paper, the formula above should have had |
24147 | 149 |
INT h. {s. h i \<le> (sub i o allocGiv)s & h i pfixGe (sub i o allocAsk)s} |
150 |
LeadsTo |
|
151 |
{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
|
152 |
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
|
153 |
looked at.*) |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
154 |
|
36866 | 155 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
156 |
\<comment> \<open>spec: preserves part\<close> |
17310 | 157 |
alloc_preserves :: "'a allocState_d program set" |
36866 | 158 |
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
|
159 |
preserves allocState_d.dummy" |
24147 | 160 |
|
36866 | 161 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
162 |
\<comment> \<open>environmental constraints\<close> |
17310 | 163 |
alloc_allowed_acts :: "'a allocState_d program set" |
36866 | 164 |
where "alloc_allowed_acts = |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
165 |
{F. AllowedActs F = |
24147 | 166 |
insert Id (UNION (preserves allocGiv) Acts)}" |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
167 |
|
36866 | 168 |
definition |
17310 | 169 |
alloc_spec :: "'a allocState_d program set" |
36866 | 170 |
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
|
171 |
alloc_allowed_acts Int alloc_preserves" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
172 |
|
67445 | 173 |
|
174 |
subsubsection \<open>Network specification\<close> |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
175 |
|
36866 | 176 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
177 |
\<comment> \<open>spec (9.1)\<close> |
17310 | 178 |
network_ask :: "'a systemState program set" |
36866 | 179 |
where "network_ask = (INT i : lessThan Nclients. |
24147 | 180 |
Increasing (ask o sub i o client) guarantees |
36866 | 181 |
((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
|
182 |
|
36866 | 183 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
184 |
\<comment> \<open>spec (9.2)\<close> |
17310 | 185 |
network_giv :: "'a systemState program set" |
36866 | 186 |
where "network_giv = (INT i : lessThan Nclients. |
24147 | 187 |
Increasing (sub i o allocGiv) |
188 |
guarantees |
|
36866 | 189 |
((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
|
190 |
|
36866 | 191 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
192 |
\<comment> \<open>spec (9.3)\<close> |
17310 | 193 |
network_rel :: "'a systemState program set" |
36866 | 194 |
where "network_rel = (INT i : lessThan Nclients. |
24147 | 195 |
Increasing (rel o sub i o client) |
196 |
guarantees |
|
36866 | 197 |
((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
|
198 |
|
36866 | 199 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
200 |
\<comment> \<open>spec: preserves part\<close> |
17310 | 201 |
network_preserves :: "'a systemState program set" |
36866 | 202 |
where "network_preserves = |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
203 |
preserves allocGiv Int |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
204 |
(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
|
205 |
preserves (ask o sub i o client))" |
24147 | 206 |
|
36866 | 207 |
definition |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64267
diff
changeset
|
208 |
\<comment> \<open>environmental constraints\<close> |
17310 | 209 |
network_allowed_acts :: "'a systemState program set" |
36866 | 210 |
where "network_allowed_acts = |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
211 |
{F. AllowedActs F = |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
212 |
insert Id |
24147 | 213 |
(UNION (preserves allocRel Int |
214 |
(INT i: lessThan Nclients. preserves(giv o sub i o client))) |
|
215 |
Acts)}" |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
216 |
|
36866 | 217 |
definition |
17310 | 218 |
network_spec :: "'a systemState program set" |
36866 | 219 |
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
|
220 |
network_rel Int network_allowed_acts Int |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
221 |
network_preserves" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
222 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
223 |
|
67445 | 224 |
subsubsection \<open>State mappings\<close> |
225 |
||
36866 | 226 |
definition |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
227 |
sysOfAlloc :: "((nat => clientState) * 'a) allocState_d => 'a systemState" |
36866 | 228 |
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
|
229 |
in (| allocGiv = allocGiv s, |
24147 | 230 |
allocAsk = allocAsk s, |
231 |
allocRel = allocRel s, |
|
232 |
client = cl, |
|
36866 | 233 |
dummy = xtr|))" |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
234 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
235 |
|
36866 | 236 |
definition |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
237 |
sysOfClient :: "(nat => clientState) * 'a allocState_d => 'a systemState" |
36866 | 238 |
where "sysOfClient = (%(cl,al). (| allocGiv = allocGiv al, |
24147 | 239 |
allocAsk = allocAsk al, |
240 |
allocRel = allocRel al, |
|
241 |
client = cl, |
|
36866 | 242 |
systemState.dummy = allocState_d.dummy al|))" |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
243 |
|
36866 | 244 |
axiomatization Alloc :: "'a allocState_d program" |
67613 | 245 |
where Alloc: "Alloc \<in> alloc_spec" |
24147 | 246 |
|
36866 | 247 |
axiomatization Client :: "'a clientState_d program" |
67613 | 248 |
where Client: "Client \<in> client_spec" |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
249 |
|
36866 | 250 |
axiomatization Network :: "'a systemState program" |
67613 | 251 |
where Network: "Network \<in> network_spec" |
36866 | 252 |
|
253 |
definition System :: "'a systemState program" |
|
60773 | 254 |
where "System = rename sysOfAlloc Alloc \<squnion> Network \<squnion> |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
255 |
(rename sysOfClient |
24147 | 256 |
(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
|
257 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
258 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
259 |
(** |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
260 |
locale System = |
24147 | 261 |
fixes |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
262 |
Alloc :: 'a allocState_d program |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
263 |
Client :: 'a clientState_d program |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
264 |
Network :: 'a systemState program |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
265 |
System :: 'a systemState program |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
266 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
267 |
assumes |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
268 |
Alloc "Alloc : alloc_spec" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
269 |
Client "Client : client_spec" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
270 |
Network "Network : network_spec" |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
271 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
272 |
defines |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
273 |
System_def |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
274 |
"System == rename sysOfAlloc Alloc |
60773 | 275 |
\<squnion> |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
276 |
Network |
60773 | 277 |
\<squnion> |
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
278 |
(rename sysOfClient |
24147 | 279 |
(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
|
280 |
**) |
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
281 |
|
21632 | 282 |
declare subset_preserves_o [THEN [2] rev_subsetD, intro] |
283 |
declare subset_preserves_o [THEN [2] rev_subsetD, simp] |
|
284 |
declare funPair_o_distrib [simp] |
|
285 |
declare Always_INT_distrib [simp] |
|
286 |
declare o_apply [simp del] |
|
287 |
||
288 |
(*For rewriting of specifications related by "guarantees"*) |
|
289 |
lemmas [simp] = |
|
290 |
rename_image_constrains |
|
291 |
rename_image_stable |
|
292 |
rename_image_increasing |
|
293 |
rename_image_invariant |
|
294 |
rename_image_Constrains |
|
295 |
rename_image_Stable |
|
296 |
rename_image_Increasing |
|
297 |
rename_image_Always |
|
298 |
rename_image_leadsTo |
|
299 |
rename_image_LeadsTo |
|
300 |
rename_preserves |
|
301 |
rename_image_preserves |
|
302 |
lift_image_preserves |
|
303 |
bij_image_INT |
|
304 |
bij_is_inj [THEN image_Int] |
|
305 |
bij_image_Collect_eq |
|
306 |
||
63146 | 307 |
ML \<open> |
21632 | 308 |
(*Splits up conjunctions & intersections: like CONJUNCTS in the HOL system*) |
24147 | 309 |
fun list_of_Int th = |
21632 | 310 |
(list_of_Int (th RS conjunct1) @ list_of_Int (th RS conjunct2)) |
311 |
handle THM _ => (list_of_Int (th RS IntD1) @ list_of_Int (th RS IntD2)) |
|
24147 | 312 |
handle THM _ => (list_of_Int (th RS @{thm INT_D})) |
21632 | 313 |
handle THM _ => (list_of_Int (th RS bspec)) |
314 |
handle THM _ => [th]; |
|
63146 | 315 |
\<close> |
21632 | 316 |
|
317 |
lemmas lessThanBspec = lessThan_iff [THEN iffD2, THEN [2] bspec] |
|
318 |
||
63146 | 319 |
attribute_setup normalized = \<open> |
24147 | 320 |
let |
321 |
fun normalized th = |
|
322 |
normalized (th RS spec |
|
323 |
handle THM _ => th RS @{thm lessThanBspec} |
|
324 |
handle THM _ => th RS bspec |
|
325 |
handle THM _ => th RS (@{thm guarantees_INT_right_iff} RS iffD1)) |
|
326 |
handle THM _ => th; |
|
21632 | 327 |
in |
61853
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents:
60773
diff
changeset
|
328 |
Scan.succeed (Thm.rule_attribute [] (K normalized)) |
21632 | 329 |
end |
63146 | 330 |
\<close> |
21632 | 331 |
|
332 |
(*** bijectivity of sysOfAlloc [MUST BE AUTOMATED] ***) |
|
63146 | 333 |
ML \<open> |
42767 | 334 |
fun record_auto_tac ctxt = |
42793 | 335 |
let val ctxt' = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
336 |
ctxt addSWrapper Record.split_wrapper |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
337 |
addsimps |
42795
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42793
diff
changeset
|
338 |
[@{thm sysOfAlloc_def}, @{thm sysOfClient_def}, |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42793
diff
changeset
|
339 |
@{thm client_map_def}, @{thm non_dummy_def}, @{thm funPair_def}, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
340 |
@{thm o_apply}, @{thm Let_def}] |
42793 | 341 |
in auto_tac ctxt' end; |
342 |
||
63146 | 343 |
\<close> |
21632 | 344 |
|
63146 | 345 |
method_setup record_auto = \<open>Scan.succeed (SIMPLE_METHOD o record_auto_tac)\<close> |
21632 | 346 |
|
347 |
lemma inj_sysOfAlloc [iff]: "inj sysOfAlloc" |
|
348 |
apply (unfold sysOfAlloc_def Let_def) |
|
349 |
apply (rule inj_onI) |
|
24147 | 350 |
apply record_auto |
21632 | 351 |
done |
352 |
||
63146 | 353 |
text\<open>We need the inverse; also having it simplifies the proof of surjectivity\<close> |
24147 | 354 |
lemma inv_sysOfAlloc_eq [simp]: "!!s. inv sysOfAlloc s = |
355 |
(| allocGiv = allocGiv s, |
|
356 |
allocAsk = allocAsk s, |
|
357 |
allocRel = allocRel s, |
|
21632 | 358 |
allocState_d.dummy = (client s, dummy s) |)" |
359 |
apply (rule inj_sysOfAlloc [THEN inv_f_eq]) |
|
24147 | 360 |
apply record_auto |
21632 | 361 |
done |
362 |
||
363 |
lemma surj_sysOfAlloc [iff]: "surj sysOfAlloc" |
|
40702 | 364 |
apply (simp add: surj_iff_all) |
24147 | 365 |
apply record_auto |
21632 | 366 |
done |
367 |
||
368 |
lemma bij_sysOfAlloc [iff]: "bij sysOfAlloc" |
|
369 |
apply (blast intro: bijI) |
|
370 |
done |
|
371 |
||
372 |
||
63146 | 373 |
subsubsection\<open>bijectivity of @{term sysOfClient}\<close> |
21632 | 374 |
|
375 |
lemma inj_sysOfClient [iff]: "inj sysOfClient" |
|
376 |
apply (unfold sysOfClient_def) |
|
377 |
apply (rule inj_onI) |
|
24147 | 378 |
apply record_auto |
21632 | 379 |
done |
380 |
||
24147 | 381 |
lemma inv_sysOfClient_eq [simp]: "!!s. inv sysOfClient s = |
382 |
(client s, |
|
383 |
(| allocGiv = allocGiv s, |
|
384 |
allocAsk = allocAsk s, |
|
385 |
allocRel = allocRel s, |
|
21632 | 386 |
allocState_d.dummy = systemState.dummy s|) )" |
387 |
apply (rule inj_sysOfClient [THEN inv_f_eq]) |
|
24147 | 388 |
apply record_auto |
21632 | 389 |
done |
390 |
||
391 |
lemma surj_sysOfClient [iff]: "surj sysOfClient" |
|
40702 | 392 |
apply (simp add: surj_iff_all) |
24147 | 393 |
apply record_auto |
21632 | 394 |
done |
395 |
||
396 |
lemma bij_sysOfClient [iff]: "bij sysOfClient" |
|
397 |
apply (blast intro: bijI) |
|
398 |
done |
|
399 |
||
400 |
||
63146 | 401 |
subsubsection\<open>bijectivity of @{term client_map}\<close> |
21632 | 402 |
|
403 |
lemma inj_client_map [iff]: "inj client_map" |
|
404 |
apply (unfold inj_on_def) |
|
24147 | 405 |
apply record_auto |
21632 | 406 |
done |
407 |
||
24147 | 408 |
lemma inv_client_map_eq [simp]: "!!s. inv client_map s = |
409 |
(%(x,y).(|giv = giv x, ask = ask x, rel = rel x, |
|
21632 | 410 |
clientState_d.dummy = y|)) s" |
411 |
apply (rule inj_client_map [THEN inv_f_eq]) |
|
24147 | 412 |
apply record_auto |
21632 | 413 |
done |
414 |
||
415 |
lemma surj_client_map [iff]: "surj client_map" |
|
40702 | 416 |
apply (simp add: surj_iff_all) |
24147 | 417 |
apply record_auto |
21632 | 418 |
done |
419 |
||
420 |
lemma bij_client_map [iff]: "bij client_map" |
|
421 |
apply (blast intro: bijI) |
|
422 |
done |
|
423 |
||
424 |
||
63146 | 425 |
text\<open>o-simprules for @{term client_map}\<close> |
21632 | 426 |
|
427 |
lemma fst_o_client_map: "fst o client_map = non_dummy" |
|
428 |
apply (unfold client_map_def) |
|
429 |
apply (rule fst_o_funPair) |
|
430 |
done |
|
431 |
||
63146 | 432 |
ML \<open>ML_Thms.bind_thms ("fst_o_client_map'", make_o_equivs @{context} @{thm fst_o_client_map})\<close> |
21632 | 433 |
declare fst_o_client_map' [simp] |
434 |
||
435 |
lemma snd_o_client_map: "snd o client_map = clientState_d.dummy" |
|
436 |
apply (unfold client_map_def) |
|
437 |
apply (rule snd_o_funPair) |
|
438 |
done |
|
439 |
||
63146 | 440 |
ML \<open>ML_Thms.bind_thms ("snd_o_client_map'", make_o_equivs @{context} @{thm snd_o_client_map})\<close> |
21632 | 441 |
declare snd_o_client_map' [simp] |
442 |
||
21710 | 443 |
|
63146 | 444 |
subsection\<open>o-simprules for @{term sysOfAlloc} [MUST BE AUTOMATED]\<close> |
21632 | 445 |
|
446 |
lemma client_o_sysOfAlloc: "client o sysOfAlloc = fst o allocState_d.dummy " |
|
24147 | 447 |
apply record_auto |
21632 | 448 |
done |
449 |
||
63146 | 450 |
ML \<open>ML_Thms.bind_thms ("client_o_sysOfAlloc'", make_o_equivs @{context} @{thm client_o_sysOfAlloc})\<close> |
21632 | 451 |
declare client_o_sysOfAlloc' [simp] |
452 |
||
453 |
lemma allocGiv_o_sysOfAlloc_eq: "allocGiv o sysOfAlloc = allocGiv" |
|
24147 | 454 |
apply record_auto |
21632 | 455 |
done |
456 |
||
63146 | 457 |
ML \<open>ML_Thms.bind_thms ("allocGiv_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocGiv_o_sysOfAlloc_eq})\<close> |
21632 | 458 |
declare allocGiv_o_sysOfAlloc_eq' [simp] |
459 |
||
460 |
lemma allocAsk_o_sysOfAlloc_eq: "allocAsk o sysOfAlloc = allocAsk" |
|
24147 | 461 |
apply record_auto |
21632 | 462 |
done |
463 |
||
63146 | 464 |
ML \<open>ML_Thms.bind_thms ("allocAsk_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocAsk_o_sysOfAlloc_eq})\<close> |
21632 | 465 |
declare allocAsk_o_sysOfAlloc_eq' [simp] |
466 |
||
467 |
lemma allocRel_o_sysOfAlloc_eq: "allocRel o sysOfAlloc = allocRel" |
|
24147 | 468 |
apply record_auto |
21632 | 469 |
done |
470 |
||
63146 | 471 |
ML \<open>ML_Thms.bind_thms ("allocRel_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocRel_o_sysOfAlloc_eq})\<close> |
21632 | 472 |
declare allocRel_o_sysOfAlloc_eq' [simp] |
473 |
||
21710 | 474 |
|
63146 | 475 |
subsection\<open>o-simprules for @{term sysOfClient} [MUST BE AUTOMATED]\<close> |
21632 | 476 |
|
477 |
lemma client_o_sysOfClient: "client o sysOfClient = fst" |
|
24147 | 478 |
apply record_auto |
21632 | 479 |
done |
480 |
||
63146 | 481 |
ML \<open>ML_Thms.bind_thms ("client_o_sysOfClient'", make_o_equivs @{context} @{thm client_o_sysOfClient})\<close> |
21632 | 482 |
declare client_o_sysOfClient' [simp] |
483 |
||
484 |
lemma allocGiv_o_sysOfClient_eq: "allocGiv o sysOfClient = allocGiv o snd " |
|
24147 | 485 |
apply record_auto |
21632 | 486 |
done |
487 |
||
63146 | 488 |
ML \<open>ML_Thms.bind_thms ("allocGiv_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocGiv_o_sysOfClient_eq})\<close> |
21632 | 489 |
declare allocGiv_o_sysOfClient_eq' [simp] |
490 |
||
491 |
lemma allocAsk_o_sysOfClient_eq: "allocAsk o sysOfClient = allocAsk o snd " |
|
24147 | 492 |
apply record_auto |
21632 | 493 |
done |
494 |
||
63146 | 495 |
ML \<open>ML_Thms.bind_thms ("allocAsk_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocAsk_o_sysOfClient_eq})\<close> |
21632 | 496 |
declare allocAsk_o_sysOfClient_eq' [simp] |
497 |
||
498 |
lemma allocRel_o_sysOfClient_eq: "allocRel o sysOfClient = allocRel o snd " |
|
24147 | 499 |
apply record_auto |
21632 | 500 |
done |
501 |
||
63146 | 502 |
ML \<open>ML_Thms.bind_thms ("allocRel_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocRel_o_sysOfClient_eq})\<close> |
21632 | 503 |
declare allocRel_o_sysOfClient_eq' [simp] |
504 |
||
505 |
lemma allocGiv_o_inv_sysOfAlloc_eq: "allocGiv o inv sysOfAlloc = allocGiv" |
|
506 |
apply (simp add: o_def) |
|
507 |
done |
|
508 |
||
63146 | 509 |
ML \<open>ML_Thms.bind_thms ("allocGiv_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocGiv_o_inv_sysOfAlloc_eq})\<close> |
21632 | 510 |
declare allocGiv_o_inv_sysOfAlloc_eq' [simp] |
511 |
||
512 |
lemma allocAsk_o_inv_sysOfAlloc_eq: "allocAsk o inv sysOfAlloc = allocAsk" |
|
513 |
apply (simp add: o_def) |
|
514 |
done |
|
515 |
||
63146 | 516 |
ML \<open>ML_Thms.bind_thms ("allocAsk_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocAsk_o_inv_sysOfAlloc_eq})\<close> |
21632 | 517 |
declare allocAsk_o_inv_sysOfAlloc_eq' [simp] |
518 |
||
519 |
lemma allocRel_o_inv_sysOfAlloc_eq: "allocRel o inv sysOfAlloc = allocRel" |
|
520 |
apply (simp add: o_def) |
|
521 |
done |
|
522 |
||
63146 | 523 |
ML \<open>ML_Thms.bind_thms ("allocRel_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocRel_o_inv_sysOfAlloc_eq})\<close> |
21632 | 524 |
declare allocRel_o_inv_sysOfAlloc_eq' [simp] |
525 |
||
24147 | 526 |
lemma rel_inv_client_map_drop_map: "(rel o inv client_map o drop_map i o inv sysOfClient) = |
21632 | 527 |
rel o sub i o client" |
528 |
apply (simp add: o_def drop_map_def) |
|
529 |
done |
|
530 |
||
63146 | 531 |
ML \<open>ML_Thms.bind_thms ("rel_inv_client_map_drop_map'", make_o_equivs @{context} @{thm rel_inv_client_map_drop_map})\<close> |
21632 | 532 |
declare rel_inv_client_map_drop_map [simp] |
533 |
||
24147 | 534 |
lemma ask_inv_client_map_drop_map: "(ask o inv client_map o drop_map i o inv sysOfClient) = |
21632 | 535 |
ask o sub i o client" |
536 |
apply (simp add: o_def drop_map_def) |
|
537 |
done |
|
538 |
||
63146 | 539 |
ML \<open>ML_Thms.bind_thms ("ask_inv_client_map_drop_map'", make_o_equivs @{context} @{thm ask_inv_client_map_drop_map})\<close> |
21632 | 540 |
declare ask_inv_client_map_drop_map [simp] |
541 |
||
542 |
||
63146 | 543 |
text\<open>Client : <unfolded specification>\<close> |
21632 | 544 |
lemmas client_spec_simps = |
545 |
client_spec_def client_increasing_def client_bounded_def |
|
546 |
client_progress_def client_allowed_acts_def client_preserves_def |
|
547 |
guarantees_Int_right |
|
548 |
||
63146 | 549 |
ML \<open> |
21632 | 550 |
val [Client_Increasing_ask, Client_Increasing_rel, |
24147 | 551 |
Client_Bounded, Client_Progress, Client_AllowedActs, |
21632 | 552 |
Client_preserves_giv, Client_preserves_dummy] = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
553 |
@{thm Client} |> simplify (@{context} addsimps @{thms client_spec_simps}) |
21632 | 554 |
|> list_of_Int; |
555 |
||
56199 | 556 |
ML_Thms.bind_thm ("Client_Increasing_ask", Client_Increasing_ask); |
557 |
ML_Thms.bind_thm ("Client_Increasing_rel", Client_Increasing_rel); |
|
558 |
ML_Thms.bind_thm ("Client_Bounded", Client_Bounded); |
|
559 |
ML_Thms.bind_thm ("Client_Progress", Client_Progress); |
|
560 |
ML_Thms.bind_thm ("Client_AllowedActs", Client_AllowedActs); |
|
561 |
ML_Thms.bind_thm ("Client_preserves_giv", Client_preserves_giv); |
|
562 |
ML_Thms.bind_thm ("Client_preserves_dummy", Client_preserves_dummy); |
|
63146 | 563 |
\<close> |
21632 | 564 |
|
565 |
declare |
|
566 |
Client_Increasing_ask [iff] |
|
567 |
Client_Increasing_rel [iff] |
|
568 |
Client_Bounded [iff] |
|
569 |
Client_preserves_giv [iff] |
|
570 |
Client_preserves_dummy [iff] |
|
571 |
||
572 |
||
63146 | 573 |
text\<open>Network : <unfolded specification>\<close> |
21632 | 574 |
lemmas network_spec_simps = |
575 |
network_spec_def network_ask_def network_giv_def |
|
576 |
network_rel_def network_allowed_acts_def network_preserves_def |
|
577 |
ball_conj_distrib |
|
578 |
||
63146 | 579 |
ML \<open> |
21632 | 580 |
val [Network_Ask, Network_Giv, Network_Rel, Network_AllowedActs, |
24147 | 581 |
Network_preserves_allocGiv, Network_preserves_rel, |
582 |
Network_preserves_ask] = |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
583 |
@{thm Network} |> simplify (@{context} addsimps @{thms network_spec_simps}) |
21632 | 584 |
|> list_of_Int; |
585 |
||
56199 | 586 |
ML_Thms.bind_thm ("Network_Ask", Network_Ask); |
587 |
ML_Thms.bind_thm ("Network_Giv", Network_Giv); |
|
588 |
ML_Thms.bind_thm ("Network_Rel", Network_Rel); |
|
589 |
ML_Thms.bind_thm ("Network_AllowedActs", Network_AllowedActs); |
|
590 |
ML_Thms.bind_thm ("Network_preserves_allocGiv", Network_preserves_allocGiv); |
|
591 |
ML_Thms.bind_thm ("Network_preserves_rel", Network_preserves_rel); |
|
592 |
ML_Thms.bind_thm ("Network_preserves_ask", Network_preserves_ask); |
|
63146 | 593 |
\<close> |
21632 | 594 |
|
595 |
declare Network_preserves_allocGiv [iff] |
|
596 |
||
597 |
declare |
|
598 |
Network_preserves_rel [simp] |
|
599 |
Network_preserves_ask [simp] |
|
600 |
||
601 |
declare |
|
602 |
Network_preserves_rel [simplified o_def, simp] |
|
603 |
Network_preserves_ask [simplified o_def, simp] |
|
604 |
||
63146 | 605 |
text\<open>Alloc : <unfolded specification>\<close> |
21632 | 606 |
lemmas alloc_spec_simps = |
607 |
alloc_spec_def alloc_increasing_def alloc_safety_def |
|
608 |
alloc_progress_def alloc_allowed_acts_def alloc_preserves_def |
|
609 |
||
63146 | 610 |
ML \<open> |
21632 | 611 |
val [Alloc_Increasing_0, Alloc_Safety, Alloc_Progress, Alloc_AllowedActs, |
24147 | 612 |
Alloc_preserves_allocRel, Alloc_preserves_allocAsk, |
613 |
Alloc_preserves_dummy] = |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
614 |
@{thm Alloc} |> simplify (@{context} addsimps @{thms alloc_spec_simps}) |
21632 | 615 |
|> list_of_Int; |
616 |
||
56199 | 617 |
ML_Thms.bind_thm ("Alloc_Increasing_0", Alloc_Increasing_0); |
618 |
ML_Thms.bind_thm ("Alloc_Safety", Alloc_Safety); |
|
619 |
ML_Thms.bind_thm ("Alloc_Progress", Alloc_Progress); |
|
620 |
ML_Thms.bind_thm ("Alloc_AllowedActs", Alloc_AllowedActs); |
|
621 |
ML_Thms.bind_thm ("Alloc_preserves_allocRel", Alloc_preserves_allocRel); |
|
622 |
ML_Thms.bind_thm ("Alloc_preserves_allocAsk", Alloc_preserves_allocAsk); |
|
623 |
ML_Thms.bind_thm ("Alloc_preserves_dummy", Alloc_preserves_dummy); |
|
63146 | 624 |
\<close> |
21632 | 625 |
|
63146 | 626 |
text\<open>Strip off the INT in the guarantees postcondition\<close> |
24147 | 627 |
|
628 |
lemmas Alloc_Increasing = Alloc_Increasing_0 [normalized] |
|
21632 | 629 |
|
630 |
declare |
|
631 |
Alloc_preserves_allocRel [iff] |
|
632 |
Alloc_preserves_allocAsk [iff] |
|
633 |
Alloc_preserves_dummy [iff] |
|
634 |
||
635 |
||
63146 | 636 |
subsection\<open>Components Lemmas [MUST BE AUTOMATED]\<close> |
21632 | 637 |
|
60773 | 638 |
lemma Network_component_System: "Network \<squnion> |
24147 | 639 |
((rename sysOfClient |
60773 | 640 |
(plam x: (lessThan Nclients). rename client_map Client)) \<squnion> |
24147 | 641 |
rename sysOfAlloc Alloc) |
21632 | 642 |
= System" |
21710 | 643 |
by (simp add: System_def Join_ac) |
21632 | 644 |
|
24147 | 645 |
lemma Client_component_System: "(rename sysOfClient |
60773 | 646 |
(plam x: (lessThan Nclients). rename client_map Client)) \<squnion> |
647 |
(Network \<squnion> rename sysOfAlloc Alloc) = System" |
|
21710 | 648 |
by (simp add: System_def Join_ac) |
21632 | 649 |
|
60773 | 650 |
lemma Alloc_component_System: "rename sysOfAlloc Alloc \<squnion> |
651 |
((rename sysOfClient (plam x: (lessThan Nclients). rename client_map Client)) \<squnion> |
|
21632 | 652 |
Network) = System" |
21710 | 653 |
by (simp add: System_def Join_ac) |
21632 | 654 |
|
655 |
declare |
|
656 |
Client_component_System [iff] |
|
657 |
Network_component_System [iff] |
|
658 |
Alloc_component_System [iff] |
|
659 |
||
21710 | 660 |
|
63146 | 661 |
text\<open>* These preservation laws should be generated automatically *\<close> |
21632 | 662 |
|
663 |
lemma Client_Allowed [simp]: "Allowed Client = preserves rel Int preserves ask" |
|
21710 | 664 |
by (auto simp add: Allowed_def Client_AllowedActs safety_prop_Acts_iff) |
21632 | 665 |
|
24147 | 666 |
lemma Network_Allowed [simp]: "Allowed Network = |
667 |
preserves allocRel Int |
|
21632 | 668 |
(INT i: lessThan Nclients. preserves(giv o sub i o client))" |
21710 | 669 |
by (auto simp add: Allowed_def Network_AllowedActs safety_prop_Acts_iff) |
21632 | 670 |
|
671 |
lemma Alloc_Allowed [simp]: "Allowed Alloc = preserves allocGiv" |
|
21710 | 672 |
by (auto simp add: Allowed_def Alloc_AllowedActs safety_prop_Acts_iff) |
21632 | 673 |
|
63146 | 674 |
text\<open>needed in \<open>rename_client_map_tac\<close>\<close> |
21632 | 675 |
lemma OK_lift_rename_Client [simp]: "OK I (%i. lift i (rename client_map Client))" |
676 |
apply (rule OK_lift_I) |
|
677 |
apply auto |
|
678 |
apply (drule_tac w1 = rel in subset_preserves_o [THEN [2] rev_subsetD]) |
|
679 |
apply (drule_tac [2] w1 = ask in subset_preserves_o [THEN [2] rev_subsetD]) |
|
680 |
apply (auto simp add: o_def split_def) |
|
681 |
done |
|
682 |
||
21710 | 683 |
lemma fst_lift_map_eq_fst [simp]: "fst (lift_map i x) i = fst x" |
684 |
apply (insert fst_o_lift_map [of i]) |
|
24147 | 685 |
apply (drule fun_cong [where x=x]) |
40702 | 686 |
apply (simp add: o_def) |
21710 | 687 |
done |
688 |
||
689 |
lemma fst_o_lift_map' [simp]: |
|
690 |
"(f \<circ> sub i \<circ> fst \<circ> lift_map i \<circ> g) = f o fst o g" |
|
24147 | 691 |
apply (subst fst_o_lift_map [symmetric]) |
692 |
apply (simp only: o_assoc) |
|
21710 | 693 |
done |
694 |
||
21632 | 695 |
|
696 |
(*The proofs of rename_Client_Increasing, rename_Client_Bounded and |
|
697 |
rename_Client_Progress are similar. All require copying out the original |
|
698 |
Client property. A forward proof can be constructed as follows: |
|
699 |
||
700 |
Client_Increasing_ask RS |
|
701 |
(bij_client_map RS rename_rename_guarantees_eq RS iffD2) |
|
702 |
RS (lift_lift_guarantees_eq RS iffD2) |
|
703 |
RS guarantees_PLam_I |
|
704 |
RS (bij_sysOfClient RS rename_rename_guarantees_eq RS iffD2) |
|
24147 | 705 |
|> simplify (simpset() addsimps [lift_image_eq_rename, o_def, split_def, |
40702 | 706 |
surj_rename]) |
21632 | 707 |
|
708 |
However, the "preserves" property remains to be discharged, and the unfolding |
|
709 |
of "o" and "sub" complicates subsequent reasoning. |
|
710 |
||
711 |
The following tactic works for all three proofs, though it certainly looks |
|
712 |
ad-hoc! |
|
713 |
*) |
|
714 |
ML |
|
63146 | 715 |
\<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
716 |
fun rename_client_map_tac ctxt = |
21632 | 717 |
EVERY [ |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
718 |
simp_tac (ctxt addsimps [@{thm rename_guarantees_eq_rename_inv}]) 1, |
60754 | 719 |
resolve_tac ctxt @{thms guarantees_PLam_I} 1, |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
56199
diff
changeset
|
720 |
assume_tac ctxt 2, |
24147 | 721 |
(*preserves: routine reasoning*) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
722 |
asm_simp_tac (ctxt addsimps [@{thm lift_preserves_sub}]) 2, |
24147 | 723 |
(*the guarantee for "lift i (rename client_map Client)" *) |
21632 | 724 |
asm_simp_tac |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
725 |
(ctxt addsimps [@{thm lift_guarantees_eq_lift_inv}, |
24147 | 726 |
@{thm rename_guarantees_eq_rename_inv}, |
40702 | 727 |
@{thm bij_imp_bij_inv}, @{thm surj_rename}, |
24147 | 728 |
@{thm inv_inv_eq}]) 1, |
52089 | 729 |
asm_simp_tac |
730 |
(ctxt addsimps [@{thm o_def}, @{thm non_dummy_def}, @{thm guarantees_Int_right}]) 1] |
|
63146 | 731 |
\<close> |
21632 | 732 |
|
63146 | 733 |
method_setup rename_client_map = \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
734 |
Scan.succeed (fn ctxt => SIMPLE_METHOD (rename_client_map_tac ctxt)) |
63146 | 735 |
\<close> |
24147 | 736 |
|
63146 | 737 |
text\<open>Lifting \<open>Client_Increasing\<close> to @{term systemState}\<close> |
67613 | 738 |
lemma rename_Client_Increasing: "i \<in> I |
739 |
==> rename sysOfClient (plam x: I. rename client_map Client) \<in> |
|
24147 | 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 |
|
67613 | 745 |
lemma preserves_sub_fst_lift_map: "[| F \<in> preserves w; i \<noteq> j |] |
746 |
==> F \<in> preserves (sub i o fst o lift_map j o funPair v w)" |
|
21632 | 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 |] |
67613 | 753 |
==> Client \<in> preserves (giv o sub i o fst o lift_map j o client_map)" |
46911 | 754 |
apply (cases "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 |
||
63146 | 779 |
text\<open>The "ok" laws, re-oriented. |
780 |
But not sure this works: theorem \<open>ok_commute\<close> is needed below\<close> |
|
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 |
|
67613 | 787 |
==> System \<in> 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 = |
|
45605 | 794 |
bij_sysOfAlloc [THEN rename_rename_guarantees_eq, THEN iffD2] |
21632 | 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, |
40702 | 801 |
simplified surj_rename o_def sub_apply |
24147 | 802 |
rename_image_Increasing bij_sysOfAlloc |
40702 | 803 |
allocGiv_o_inv_sysOfAlloc_eq'] |
21632 | 804 |
|
24147 | 805 |
lemma System_Increasing_allocGiv: |
67613 | 806 |
"i < Nclients \<Longrightarrow> System \<in> Increasing (sub i o allocGiv)" |
21632 | 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 |
|
63146 | 814 |
ML \<open> |
56199 | 815 |
ML_Thms.bind_thms ("System_Increasing'", list_of_Int @{thm System_Increasing}) |
63146 | 816 |
\<close> |
21632 | 817 |
|
818 |
declare System_Increasing' [intro!] |
|
819 |
||
63146 | 820 |
text\<open>Follows consequences. |
21632 | 821 |
The "Always (INT ...) formulation expresses the general safety property |
63146 | 822 |
and allows it to be combined using \<open>Always_Int_rule\<close> below.\<close> |
21632 | 823 |
|
24147 | 824 |
lemma System_Follows_rel: |
67613 | 825 |
"i < Nclients ==> System \<in> ((sub i o allocRel) Fols (rel o sub i o client))" |
21632 | 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: |
67613 | 831 |
"i < Nclients ==> System \<in> ((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: |
67613 | 837 |
"i < Nclients ==> System \<in> (giv o sub i o client) Fols (sub i o allocGiv)" |
21632 | 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 |
|
67613 | 845 |
lemma Always_giv_le_allocGiv: "System \<in> 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 |
|
67613 | 852 |
lemma Always_allocAsk_le_ask: "System \<in> 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 |
|
67613 | 859 |
lemma Always_allocRel_le_rel: "System \<in> 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 |
|
63146 | 864 |
subsection\<open>Proof of the safety property (1)\<close> |
21632 | 865 |
|
63146 | 866 |
text\<open>safety (1), step 1 is \<open>System_Follows_rel\<close>\<close> |
21632 | 867 |
|
63146 | 868 |
text\<open>safety (1), step 2\<close> |
21632 | 869 |
(* i < Nclients ==> System : Increasing (sub i o allocRel) *) |
45605 | 870 |
lemmas System_Increasing_allocRel = System_Follows_rel [THEN Follows_Increasing1] |
21632 | 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 |
||
63146 | 876 |
text\<open>safety (1), step 3\<close> |
24147 | 877 |
lemma System_sum_bounded: |
67613 | 878 |
"System \<in> Always {s. (\<Sum>i \<in> lessThan Nclients. (tokens o sub i o allocGiv) s) |
21710 | 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]) |
40702 | 882 |
apply (simp add: o_def) |
24147 | 883 |
apply (erule component_guaranteesD) |
21710 | 884 |
apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def]) |
21632 | 885 |
done |
886 |
||
63146 | 887 |
text\<open>Follows reasoning\<close> |
21632 | 888 |
|
67613 | 889 |
lemma Always_tokens_giv_le_allocGiv: "System \<in> Always (INT i: lessThan Nclients. |
24147 | 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 |
||
67613 | 896 |
lemma Always_tokens_allocRel_le_rel: "System \<in> Always (INT i: lessThan Nclients. |
24147 | 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 |
||
63146 | 903 |
text\<open>safety (1), step 4 (final result!)\<close> |
67613 | 904 |
theorem System_safety: "System \<in> system_safety" |
21632 | 905 |
apply (unfold system_safety_def) |
63146 | 906 |
apply (tactic \<open>resolve_tac @{context} [Always_Int_rule [@{thm System_sum_bounded}, |
39159 | 907 |
@{thm Always_tokens_giv_le_allocGiv}, @{thm Always_tokens_allocRel_le_rel}] RS |
63146 | 908 |
@{thm Always_weaken}] 1\<close>) |
21632 | 909 |
apply auto |
64267 | 910 |
apply (rule sum_fun_mono [THEN order_trans]) |
21632 | 911 |
apply (drule_tac [2] order_trans) |
64267 | 912 |
apply (rule_tac [2] add_le_mono [OF order_refl sum_fun_mono]) |
21632 | 913 |
prefer 3 apply assumption |
914 |
apply auto |
|
915 |
done |
|
916 |
||
63146 | 917 |
subsection \<open>Proof of the progress property (2)\<close> |
21632 | 918 |
|
63146 | 919 |
text\<open>progress (2), step 1 is \<open>System_Follows_ask\<close> and |
920 |
\<open>System_Follows_rel\<close>\<close> |
|
21632 | 921 |
|
63146 | 922 |
text\<open>progress (2), step 2; see also \<open>System_Increasing_allocRel\<close>\<close> |
21632 | 923 |
(* i < Nclients ==> System : Increasing (sub i o allocAsk) *) |
45605 | 924 |
lemmas System_Increasing_allocAsk = System_Follows_ask [THEN Follows_Increasing1] |
21632 | 925 |
|
63146 | 926 |
text\<open>progress (2), step 3: lifting \<open>Client_Bounded\<close> to systemState\<close> |
67613 | 927 |
lemma rename_Client_Bounded: "i \<in> I |
928 |
==> rename sysOfClient (plam x: I. rename client_map Client) \<in> |
|
24147 | 929 |
UNIV guarantees |
67613 | 930 |
Always {s. \<forall>elt \<in> 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 |
67613 | 935 |
==> System \<in> Always |
936 |
{s. \<forall>elt \<in> 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 |
||
67613 | 941 |
lemma Collect_all_imp_eq: "{x. \<forall>y. P y \<longrightarrow> Q x y} = (INT y: {y. P y}. {x. Q x y})" |
21632 | 942 |
apply blast |
943 |
done |
|
944 |
||
63146 | 945 |
text\<open>progress (2), step 4\<close> |
67717 | 946 |
lemma System_Bounded_allocAsk: "System \<in> Always {s. \<forall>i<Nclients. |
67613 | 947 |
\<forall>elt \<in> set ((sub i o allocAsk) s). elt \<le> NbT}" |
21632 | 948 |
apply (auto simp add: Collect_all_imp_eq) |
63146 | 949 |
apply (tactic \<open>resolve_tac @{context} [Always_Int_rule [@{thm Always_allocAsk_le_ask}, |
950 |
@{thm System_Bounded_ask}] RS @{thm Always_weaken}] 1\<close>) |
|
21632 | 951 |
apply (auto dest: set_mono) |
952 |
done |
|
953 |
||
63146 | 954 |
text\<open>progress (2), step 5 is \<open>System_Increasing_allocGiv\<close>\<close> |
21632 | 955 |
|
63146 | 956 |
text\<open>progress (2), step 6\<close> |
21632 | 957 |
(* i < Nclients ==> System : Increasing (giv o sub i o client) *) |
45605 | 958 |
lemmas System_Increasing_giv = System_Follows_allocGiv [THEN Follows_Increasing1] |
21632 | 959 |
|
960 |
||
67613 | 961 |
lemma rename_Client_Progress: "i \<in> I |
24147 | 962 |
==> rename sysOfClient (plam x: I. rename client_map Client) |
67613 | 963 |
\<in> Increasing (giv o sub i o client) |
24147 | 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 |
||
63146 | 973 |
text\<open>progress (2), step 7\<close> |
24147 | 974 |
lemma System_Client_Progress: |
67613 | 975 |
"System \<in> (INT i : (lessThan Nclients). |
24147 | 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 |
67613 | 1001 |
==> System \<in> {s. h \<le> (sub i o allocGiv) s & |
24147 | 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) |
|
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
44871
diff
changeset
|
1007 |
apply (rule_tac k1 = h and x1 = "(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 |
||
63146 | 1014 |
text\<open>progress (2), step 8: Client i's "release" action is visible system-wide\<close> |
24147 | 1015 |
lemma System_Alloc_Client_Progress: "i < Nclients |
67613 | 1016 |
==> System \<in> {s. h \<le> (sub i o allocGiv) s & |
24147 | 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 |
||
63146 | 1032 |
text\<open>Lifting \<open>Alloc_Progress\<close> up to the level of systemState\<close> |
21632 | 1033 |
|
63146 | 1034 |
text\<open>progress (2), step 9\<close> |
24147 | 1035 |
lemma System_Alloc_Progress: |
67613 | 1036 |
"System \<in> (INT i : (lessThan Nclients). |
24147 | 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) |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61954
diff
changeset
|
1042 |
apply (drule 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 |
||
63146 | 1050 |
text\<open>progress (2), step 10 (final result!)\<close> |
67613 | 1051 |
lemma System_Progress: "System \<in> system_progress" |
21632 | 1052 |
apply (unfold system_progress_def) |
1053 |
apply (cut_tac System_Alloc_Progress) |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61954
diff
changeset
|
1054 |
apply auto |
21632 | 1055 |
apply (blast intro: LeadsTo_Trans |
1056 |
System_Follows_allocGiv [THEN Follows_LeadsTo_pfixLe] |
|
1057 |
System_Follows_ask [THEN Follows_LeadsTo]) |
|
1058 |
done |
|
1059 |
||
1060 |
||
67613 | 1061 |
theorem System_correct: "System \<in> system_spec" |
21632 | 1062 |
apply (unfold system_spec_def) |
1063 |
apply (blast intro: System_safety System_Progress) |
|
1064 |
done |
|
1065 |
||
1066 |
||
63146 | 1067 |
text\<open>Some obsolete lemmas\<close> |
21632 | 1068 |
|
24147 | 1069 |
lemma non_dummy_eq_o_funPair: "non_dummy = (% (g,a,r). (| giv = g, ask = a, rel = r |)) o |
21632 | 1070 |
(funPair giv (funPair ask rel))" |
1071 |
apply (rule ext) |
|
1072 |
apply (auto simp add: o_def non_dummy_def) |
|
1073 |
done |
|
1074 |
||
24147 | 1075 |
lemma preserves_non_dummy_eq: "(preserves non_dummy) = |
21632 | 1076 |
(preserves rel Int preserves ask Int preserves giv)" |
1077 |
apply (simp add: non_dummy_eq_o_funPair) |
|
1078 |
apply auto |
|
1079 |
apply (drule_tac w1 = rel in subset_preserves_o [THEN [2] rev_subsetD]) |
|
1080 |
apply (drule_tac [2] w1 = ask in subset_preserves_o [THEN [2] rev_subsetD]) |
|
1081 |
apply (drule_tac [3] w1 = giv in subset_preserves_o [THEN [2] rev_subsetD]) |
|
1082 |
apply (auto simp add: o_def) |
|
1083 |
done |
|
1084 |
||
63146 | 1085 |
text\<open>Could go to Extend.ML\<close> |
67613 | 1086 |
lemma bij_fst_inv_inv_eq: "bij f \<Longrightarrow> fst (inv (%(x, u). inv f x) z) = f z" |
21632 | 1087 |
apply (rule fst_inv_equalityI) |
59807 | 1088 |
apply (rule_tac f = "%z. (f z, h z)" for h in surjI) |
21632 | 1089 |
apply (simp add: bij_is_inj inv_f_f) |
1090 |
apply (simp add: bij_is_surj surj_f_inv_f) |
|
1091 |
done |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
1092 |
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
1093 |
end |