author | wenzelm |
Fri, 31 Jul 2009 11:34:14 +0200 | |
changeset 32299 | 5f33ce0ed21f |
parent 26289 | 9d2c375e242b |
child 32960 | 69916a850301 |
permissions | -rw-r--r-- |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
1 |
(* Title: ZF/UNITY/ClientImpl.thy |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
3 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
4 |
Copyright 2002 University of Cambridge |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
5 |
|
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
6 |
Distributed Resource Management System: Client Implementation |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
7 |
*) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
8 |
|
16417 | 9 |
theory ClientImpl imports AllocBase Guar begin |
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
10 |
|
24892 | 11 |
abbreviation "ask == Var(Nil)" (* input history: tokens requested *) |
12 |
abbreviation "giv == Var([0])" (* output history: tokens granted *) |
|
13 |
abbreviation "rel == Var([1])" (* input history: tokens released *) |
|
14 |
abbreviation "tok == Var([2])" (* the number of available tokens *) |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
15 |
|
14061 | 16 |
axioms |
17 |
type_assumes: |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
18 |
"type_of(ask) = list(tokbag) & type_of(giv) = list(tokbag) & |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
19 |
type_of(rel) = list(tokbag) & type_of(tok) = nat" |
14061 | 20 |
default_val_assumes: |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
21 |
"default_val(ask) = Nil & default_val(giv) = Nil & |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
22 |
default_val(rel) = Nil & default_val(tok) = 0" |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
23 |
|
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
24 |
|
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
25 |
(*Array indexing is translated to list indexing as A[n] == nth(n-1,A). *) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
26 |
|
24893 | 27 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
28 |
(** Release some client_tokens **) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
29 |
"client_rel_act == |
14061 | 30 |
{<s,t> \<in> state*state. |
31 |
\<exists>nrel \<in> nat. nrel = length(s`rel) & |
|
14057 | 32 |
t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) & |
33 |
nrel < length(s`giv) & |
|
14061 | 34 |
nth(nrel, s`ask) \<le> nth(nrel, s`giv)}" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
35 |
|
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
36 |
(** Choose a new token requirement **) |
14057 | 37 |
(** Including t=s suppresses fairness, allowing the non-trivial part |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
38 |
of the action to be ignored **) |
24893 | 39 |
definition |
40 |
"client_tok_act == {<s,t> \<in> state*state. t=s | |
|
14057 | 41 |
t = s(tok:=succ(s`tok mod NbT))}" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
42 |
|
24893 | 43 |
definition |
14061 | 44 |
"client_ask_act == {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
45 |
|
24893 | 46 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
47 |
"client_prog == |
14061 | 48 |
mk_program({s \<in> state. s`tok \<le> NbT & s`giv = Nil & |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
49 |
s`ask = Nil & s`rel = Nil}, |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
50 |
{client_rel_act, client_tok_act, client_ask_act}, |
14061 | 51 |
\<Union>G \<in> preserves(lift(rel)) Int |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
52 |
preserves(lift(ask)) Int |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
53 |
preserves(lift(tok)). Acts(G))" |
14061 | 54 |
|
55 |
||
56 |
declare type_assumes [simp] default_val_assumes [simp] |
|
57 |
(* This part should be automated *) |
|
58 |
||
59 |
lemma ask_value_type [simp,TC]: "s \<in> state ==> s`ask \<in> list(nat)" |
|
60 |
apply (unfold state_def) |
|
61 |
apply (drule_tac a = ask in apply_type, auto) |
|
62 |
done |
|
63 |
||
64 |
lemma giv_value_type [simp,TC]: "s \<in> state ==> s`giv \<in> list(nat)" |
|
65 |
apply (unfold state_def) |
|
66 |
apply (drule_tac a = giv in apply_type, auto) |
|
67 |
done |
|
68 |
||
69 |
lemma rel_value_type [simp,TC]: "s \<in> state ==> s`rel \<in> list(nat)" |
|
70 |
apply (unfold state_def) |
|
71 |
apply (drule_tac a = rel in apply_type, auto) |
|
72 |
done |
|
73 |
||
74 |
lemma tok_value_type [simp,TC]: "s \<in> state ==> s`tok \<in> nat" |
|
75 |
apply (unfold state_def) |
|
76 |
apply (drule_tac a = tok in apply_type, auto) |
|
77 |
done |
|
78 |
||
79 |
(** The Client Program **) |
|
80 |
||
81 |
lemma client_type [simp,TC]: "client_prog \<in> program" |
|
82 |
apply (unfold client_prog_def) |
|
83 |
apply (simp (no_asm)) |
|
84 |
done |
|
85 |
||
86 |
declare client_prog_def [THEN def_prg_Init, simp] |
|
87 |
declare client_prog_def [THEN def_prg_AllowedActs, simp] |
|
24051
896fb015079c
replaced program_defs_ref by proper context data (via attribute "program");
wenzelm
parents:
16417
diff
changeset
|
88 |
declare client_prog_def [program] |
14061 | 89 |
|
90 |
declare client_rel_act_def [THEN def_act_simp, simp] |
|
91 |
declare client_tok_act_def [THEN def_act_simp, simp] |
|
92 |
declare client_ask_act_def [THEN def_act_simp, simp] |
|
93 |
||
94 |
lemma client_prog_ok_iff: |
|
95 |
"\<forall>G \<in> program. (client_prog ok G) <-> |
|
96 |
(G \<in> preserves(lift(rel)) & G \<in> preserves(lift(ask)) & |
|
97 |
G \<in> preserves(lift(tok)) & client_prog \<in> Allowed(G))" |
|
98 |
by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed]) |
|
99 |
||
100 |
lemma client_prog_preserves: |
|
101 |
"client_prog:(\<Inter>x \<in> var-{ask, rel, tok}. preserves(lift(x)))" |
|
102 |
apply (rule Inter_var_DiffI, force) |
|
103 |
apply (rule ballI) |
|
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
15634
diff
changeset
|
104 |
apply (rule preservesI, safety, auto) |
14061 | 105 |
done |
106 |
||
107 |
||
108 |
lemma preserves_lift_imp_stable: |
|
109 |
"G \<in> preserves(lift(ff)) ==> G \<in> stable({s \<in> state. P(s`ff)})"; |
|
110 |
apply (drule preserves_imp_stable) |
|
111 |
apply (simp add: lift_def) |
|
112 |
done |
|
113 |
||
114 |
lemma preserves_imp_prefix: |
|
115 |
"G \<in> preserves(lift(ff)) |
|
116 |
==> G \<in> stable({s \<in> state. \<langle>k, s`ff\<rangle> \<in> prefix(nat)})"; |
|
117 |
by (erule preserves_lift_imp_stable) |
|
118 |
||
119 |
(*Safety property 1: ask, rel are increasing: (24) *) |
|
120 |
lemma client_prog_Increasing_ask_rel: |
|
121 |
"client_prog: program guarantees Incr(lift(ask)) Int Incr(lift(rel))" |
|
122 |
apply (unfold guar_def) |
|
123 |
apply (auto intro!: increasing_imp_Increasing |
|
26289 | 124 |
simp add: client_prog_ok_iff Increasing.increasing_def preserves_imp_prefix) |
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
15634
diff
changeset
|
125 |
apply (safety, force, force)+ |
14061 | 126 |
done |
127 |
||
128 |
declare nth_append [simp] append_one_prefix [simp] |
|
129 |
||
130 |
lemma NbT_pos2: "0<NbT" |
|
131 |
apply (cut_tac NbT_pos) |
|
132 |
apply (rule Ord_0_lt, auto) |
|
133 |
done |
|
134 |
||
135 |
(*Safety property 2: the client never requests too many tokens. |
|
136 |
With no Substitution Axiom, we must prove the two invariants simultaneously. *) |
|
137 |
||
138 |
lemma ask_Bounded_lemma: |
|
139 |
"[| client_prog ok G; G \<in> program |] |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
140 |
==> client_prog \<squnion> G \<in> |
14061 | 141 |
Always({s \<in> state. s`tok \<le> NbT} Int |
142 |
{s \<in> state. \<forall>elt \<in> set_of_list(s`ask). elt \<le> NbT})" |
|
143 |
apply (rotate_tac -1) |
|
144 |
apply (auto simp add: client_prog_ok_iff) |
|
145 |
apply (rule invariantI [THEN stable_Join_Always2], force) |
|
146 |
prefer 2 |
|
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
15634
diff
changeset
|
147 |
apply (fast intro: stable_Int preserves_lift_imp_stable, safety) |
14061 | 148 |
apply (auto dest: ActsD) |
149 |
apply (cut_tac NbT_pos) |
|
150 |
apply (rule NbT_pos2 [THEN mod_less_divisor]) |
|
151 |
apply (auto dest: ActsD preserves_imp_eq simp add: set_of_list_append) |
|
152 |
done |
|
153 |
||
154 |
(* Export version, with no mention of tok in the postcondition, but |
|
155 |
unfortunately tok must be declared local.*) |
|
156 |
lemma client_prog_ask_Bounded: |
|
157 |
"client_prog \<in> program guarantees |
|
158 |
Always({s \<in> state. \<forall>elt \<in> set_of_list(s`ask). elt \<le> NbT})" |
|
159 |
apply (rule guaranteesI) |
|
160 |
apply (erule ask_Bounded_lemma [THEN Always_weaken], auto) |
|
161 |
done |
|
162 |
||
163 |
(*** Towards proving the liveness property ***) |
|
164 |
||
165 |
lemma client_prog_stable_rel_le_giv: |
|
166 |
"client_prog \<in> stable({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})" |
|
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
15634
diff
changeset
|
167 |
by (safety, auto) |
14061 | 168 |
|
169 |
lemma client_prog_Join_Stable_rel_le_giv: |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
170 |
"[| client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel)) |] |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
171 |
==> client_prog \<squnion> G \<in> Stable({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})" |
14061 | 172 |
apply (rule client_prog_stable_rel_le_giv [THEN Increasing_preserves_Stable]) |
173 |
apply (auto simp add: lift_def) |
|
174 |
done |
|
175 |
||
176 |
lemma client_prog_Join_Always_rel_le_giv: |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
177 |
"[| client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel)) |] |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
178 |
==> client_prog \<squnion> G \<in> Always({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})" |
14061 | 179 |
by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv) |
180 |
||
181 |
lemma def_act_eq: |
|
182 |
"A == {<s, t> \<in> state*state. P(s, t)} ==> A={<s, t> \<in> state*state. P(s, t)}" |
|
183 |
by auto |
|
184 |
||
185 |
lemma act_subset: "A={<s,t> \<in> state*state. P(s, t)} ==> A<=state*state" |
|
186 |
by auto |
|
187 |
||
188 |
lemma transient_lemma: |
|
189 |
"client_prog \<in> |
|
190 |
transient({s \<in> state. s`rel = k & <k, h> \<in> strict_prefix(nat) |
|
191 |
& <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask})" |
|
192 |
apply (rule_tac act = client_rel_act in transientI) |
|
193 |
apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts]) |
|
194 |
apply (simp (no_asm) add: client_rel_act_def [THEN def_act_eq, THEN act_subset]) |
|
195 |
apply (auto simp add: client_prog_def [THEN def_prg_Acts] domain_def) |
|
196 |
apply (rule ReplaceI) |
|
197 |
apply (rule_tac x = "x (rel:= x`rel @ [nth (length (x`rel), x`giv) ]) " in exI) |
|
198 |
apply (auto intro!: state_update_type app_type length_type nth_type, auto) |
|
199 |
apply (blast intro: lt_trans2 prefix_length_le strict_prefix_length_lt) |
|
200 |
apply (blast intro: lt_trans2 prefix_length_le strict_prefix_length_lt) |
|
201 |
apply (simp (no_asm_use) add: gen_prefix_iff_nth) |
|
202 |
apply (subgoal_tac "h \<in> list(nat)") |
|
203 |
apply (simp_all (no_asm_simp) add: prefix_type [THEN subsetD, THEN SigmaD1]) |
|
204 |
apply (auto simp add: prefix_def Ge_def) |
|
205 |
apply (drule strict_prefix_length_lt) |
|
206 |
apply (drule_tac x = "length (x`rel) " in spec) |
|
207 |
apply auto |
|
208 |
apply (simp (no_asm_use) add: gen_prefix_iff_nth) |
|
209 |
apply (auto simp add: id_def lam_def) |
|
210 |
done |
|
211 |
||
212 |
lemma strict_prefix_is_prefix: |
|
213 |
"<xs, ys> \<in> strict_prefix(A) <-> <xs, ys> \<in> prefix(A) & xs\<noteq>ys" |
|
214 |
apply (unfold strict_prefix_def id_def lam_def) |
|
215 |
apply (auto dest: prefix_type [THEN subsetD]) |
|
216 |
done |
|
217 |
||
218 |
lemma induct_lemma: |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
219 |
"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |] |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
220 |
==> client_prog \<squnion> G \<in> |
14061 | 221 |
{s \<in> state. s`rel = k & <k,h> \<in> strict_prefix(nat) |
222 |
& <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask} |
|
223 |
LeadsTo {s \<in> state. <k, s`rel> \<in> strict_prefix(nat) |
|
224 |
& <s`rel, s`giv> \<in> prefix(nat) & |
|
225 |
<h, s`giv> \<in> prefix(nat) & |
|
226 |
h pfixGe s`ask}" |
|
227 |
apply (rule single_LeadsTo_I) |
|
228 |
prefer 2 apply simp |
|
229 |
apply (frule client_prog_Increasing_ask_rel [THEN guaranteesD]) |
|
230 |
apply (rotate_tac [3] 2) |
|
231 |
apply (auto simp add: client_prog_ok_iff) |
|
232 |
apply (rule transient_lemma [THEN Join_transient_I1, THEN transient_imp_leadsTo, THEN leadsTo_imp_LeadsTo, THEN PSP_Stable, THEN LeadsTo_weaken]) |
|
233 |
apply (rule Stable_Int [THEN Stable_Int, THEN Stable_Int]) |
|
234 |
apply (erule_tac f = "lift (giv) " and a = "s`giv" in Increasing_imp_Stable) |
|
235 |
apply (simp (no_asm_simp)) |
|
236 |
apply (erule_tac f = "lift (ask) " and a = "s`ask" in Increasing_imp_Stable) |
|
237 |
apply (simp (no_asm_simp)) |
|
238 |
apply (erule_tac f = "lift (rel) " and a = "s`rel" in Increasing_imp_Stable) |
|
239 |
apply (simp (no_asm_simp)) |
|
240 |
apply (erule client_prog_Join_Stable_rel_le_giv, blast, simp_all) |
|
241 |
prefer 2 |
|
242 |
apply (blast intro: sym strict_prefix_is_prefix [THEN iffD2] prefix_trans prefix_imp_pfixGe pfixGe_trans) |
|
243 |
apply (auto intro: strict_prefix_is_prefix [THEN iffD1, THEN conjunct1] |
|
244 |
prefix_trans) |
|
245 |
done |
|
246 |
||
247 |
lemma rel_progress_lemma: |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
248 |
"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |] |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
249 |
==> client_prog \<squnion> G \<in> |
14061 | 250 |
{s \<in> state. <s`rel, h> \<in> strict_prefix(nat) |
251 |
& <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask} |
|
252 |
LeadsTo {s \<in> state. <h, s`rel> \<in> prefix(nat)}" |
|
253 |
apply (rule_tac f = "\<lambda>x \<in> state. length(h) #- length(x`rel)" |
|
254 |
in LessThan_induct) |
|
255 |
apply (auto simp add: vimage_def) |
|
256 |
prefer 2 apply (force simp add: lam_def) |
|
257 |
apply (rule single_LeadsTo_I) |
|
258 |
prefer 2 apply simp |
|
259 |
apply (subgoal_tac "h \<in> list(nat)") |
|
260 |
prefer 2 apply (blast dest: prefix_type [THEN subsetD]) |
|
261 |
apply (rule induct_lemma [THEN LeadsTo_weaken]) |
|
262 |
apply (simp add: length_type lam_def) |
|
263 |
apply (auto intro: strict_prefix_is_prefix [THEN iffD2] |
|
264 |
dest: common_prefix_linear prefix_type [THEN subsetD]) |
|
265 |
apply (erule swap) |
|
266 |
apply (rule imageI) |
|
267 |
apply (force dest!: simp add: lam_def) |
|
268 |
apply (simp add: length_type lam_def, clarify) |
|
269 |
apply (drule strict_prefix_length_lt)+ |
|
270 |
apply (drule less_imp_succ_add, simp)+ |
|
271 |
apply clarify |
|
272 |
apply simp |
|
273 |
apply (erule diff_le_self [THEN ltD]) |
|
274 |
done |
|
275 |
||
276 |
lemma progress_lemma: |
|
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
277 |
"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |] |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
278 |
==> client_prog \<squnion> G |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
279 |
\<in> {s \<in> state. <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask} |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
280 |
LeadsTo {s \<in> state. <h, s`rel> \<in> prefix(nat)}" |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
281 |
apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI], |
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
282 |
assumption) |
14061 | 283 |
apply (force simp add: client_prog_ok_iff) |
284 |
apply (rule LeadsTo_weaken_L) |
|
285 |
apply (rule LeadsTo_Un [OF rel_progress_lemma |
|
286 |
subset_refl [THEN subset_imp_LeadsTo]]) |
|
287 |
apply (auto intro: strict_prefix_is_prefix [THEN iffD2] |
|
288 |
dest: common_prefix_linear prefix_type [THEN subsetD]) |
|
289 |
done |
|
290 |
||
291 |
(*Progress property: all tokens that are given will be released*) |
|
292 |
lemma client_prog_progress: |
|
293 |
"client_prog \<in> Incr(lift(giv)) guarantees |
|
294 |
(\<Inter>h \<in> list(nat). {s \<in> state. <h, s`giv> \<in> prefix(nat) & |
|
295 |
h pfixGe s`ask} LeadsTo {s \<in> state. <h, s`rel> \<in> prefix(nat)})" |
|
296 |
apply (rule guaranteesI) |
|
297 |
apply (blast intro: progress_lemma, auto) |
|
298 |
done |
|
299 |
||
300 |
lemma client_prog_Allowed: |
|
301 |
"Allowed(client_prog) = |
|
302 |
preserves(lift(rel)) Int preserves(lift(ask)) Int preserves(lift(tok))" |
|
303 |
apply (cut_tac v = "lift (ask)" in preserves_type) |
|
304 |
apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed] |
|
305 |
cons_Int_distrib safety_prop_Acts_iff) |
|
306 |
done |
|
307 |
||
14072
f932be305381
Conversion of UNITY/Distributor to Isar script. General tidy-up.
paulson
parents:
14061
diff
changeset
|
308 |
end |