1(*  Title:      ZF/UNITY/AllocImpl.thy
2    Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
3    Copyright   2002  University of Cambridge
4
5Single-client allocator implementation.
6Charpentier and Chandy, section 7 (page 17).
7*)
8
9theory AllocImpl imports ClientImpl begin
10
11abbreviation
12  NbR :: i            (*number of consumed messages*)  where
13  "NbR == Var([succ(2)])"
14
15abbreviation
16  available_tok :: i  (*number of free tokens (T in paper)*)  where
17  "available_tok == Var([succ(succ(2))])"
18
19axiomatization where
20  alloc_type_assumes [simp]:
21  "type_of(NbR) = nat & type_of(available_tok)=nat" and
22
23  alloc_default_val_assumes [simp]:
24  "default_val(NbR)  = 0 & default_val(available_tok)=0"
25
26definition
27  "alloc_giv_act ==
28       {<s, t> \<in> state*state.
29        \<exists>k. k = length(s`giv) &
30            t = s(giv := s`giv @ [nth(k, s`ask)],
31                  available_tok := s`available_tok #- nth(k, s`ask)) &
32            k < length(s`ask) & nth(k, s`ask) \<le> s`available_tok}"
33
34definition
35  "alloc_rel_act ==
36       {<s, t> \<in> state*state.
37        t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
38              NbR := succ(s`NbR)) &
39        s`NbR < length(s`rel)}"
40
41definition
42  (*The initial condition s`giv=[] is missing from the
43    original definition: S. O. Ehmety *)
44  "alloc_prog ==
45       mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil},
46                  {alloc_giv_act, alloc_rel_act},
47                  \<Union>G \<in> preserves(lift(available_tok)) \<inter>
48                        preserves(lift(NbR)) \<inter>
49                        preserves(lift(giv)). Acts(G))"
50
51
52lemma available_tok_value_type [simp,TC]: "s\<in>state ==> s`available_tok \<in> nat"
53apply (unfold state_def)
54apply (drule_tac a = available_tok in apply_type, auto)
55done
56
57lemma NbR_value_type [simp,TC]: "s\<in>state ==> s`NbR \<in> nat"
58apply (unfold state_def)
59apply (drule_tac a = NbR in apply_type, auto)
60done
61
62(** The Alloc Program **)
63
64lemma alloc_prog_type [simp,TC]: "alloc_prog \<in> program"
65by (simp add: alloc_prog_def)
66
67declare alloc_prog_def [THEN def_prg_Init, simp]
68declare alloc_prog_def [THEN def_prg_AllowedActs, simp]
69declare alloc_prog_def [program]
70
71declare  alloc_giv_act_def [THEN def_act_simp, simp]
72declare  alloc_rel_act_def [THEN def_act_simp, simp]
73
74
75lemma alloc_prog_ok_iff:
76"\<forall>G \<in> program. (alloc_prog ok G) \<longleftrightarrow>
77     (G \<in> preserves(lift(giv)) & G \<in> preserves(lift(available_tok)) &
78       G \<in> preserves(lift(NbR)) &  alloc_prog \<in> Allowed(G))"
79by (auto simp add: ok_iff_Allowed alloc_prog_def [THEN def_prg_Allowed])
80
81
82lemma alloc_prog_preserves:
83    "alloc_prog \<in> (\<Inter>x \<in> var-{giv, available_tok, NbR}. preserves(lift(x)))"
84apply (rule Inter_var_DiffI, force)
85apply (rule ballI)
86apply (rule preservesI, safety)
87done
88
89(* As a special case of the rule above *)
90
91lemma alloc_prog_preserves_rel_ask_tok:
92    "alloc_prog \<in>
93       preserves(lift(rel)) \<inter> preserves(lift(ask)) \<inter> preserves(lift(tok))"
94apply auto
95apply (insert alloc_prog_preserves)
96apply (drule_tac [3] x = tok in Inter_var_DiffD)
97apply (drule_tac [2] x = ask in Inter_var_DiffD)
98apply (drule_tac x = rel in Inter_var_DiffD, auto)
99done
100
101lemma alloc_prog_Allowed:
102"Allowed(alloc_prog) =
103  preserves(lift(giv)) \<inter> preserves(lift(available_tok)) \<inter> preserves(lift(NbR))"
104apply (cut_tac v="lift(giv)" in preserves_type)
105apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed]
106                      cons_Int_distrib safety_prop_Acts_iff)
107done
108
109(* In particular we have *)
110lemma alloc_prog_ok_client_prog: "alloc_prog ok client_prog"
111apply (auto simp add: ok_iff_Allowed)
112apply (cut_tac alloc_prog_preserves)
113apply (cut_tac [2] client_prog_preserves)
114apply (auto simp add: alloc_prog_Allowed client_prog_Allowed)
115apply (drule_tac [6] B = "preserves (lift (NbR))" in InterD)
116apply (drule_tac [5] B = "preserves (lift (available_tok))" in InterD)
117apply (drule_tac [4] B = "preserves (lift (giv))" in InterD)
118apply (drule_tac [3] B = "preserves (lift (tok))" in InterD)
119apply (drule_tac [2] B = "preserves (lift (ask))" in InterD)
120apply (drule_tac B = "preserves (lift (rel))" in InterD)
121apply auto
122done
123
124(** Safety property: (28) **)
125lemma alloc_prog_Increasing_giv: "alloc_prog \<in> program guarantees Incr(lift(giv))"
126apply (auto intro!: increasing_imp_Increasing simp add: guar_def
127  Increasing.increasing_def alloc_prog_ok_iff alloc_prog_Allowed, safety+)
128apply (auto dest: ActsD)
129apply (drule_tac f = "lift (giv) " in preserves_imp_eq)
130apply auto
131done
132
133lemma giv_Bounded_lamma1:
134"alloc_prog \<in> stable({s\<in>state. s`NbR \<le> length(s`rel)} \<inter>
135                     {s\<in>state. s`available_tok #+ tokens(s`giv) =
136                                 NbT #+ tokens(take(s`NbR, s`rel))})"
137apply safety
138apply auto
139apply (simp add: diff_add_0 add_commute diff_add_inverse add_assoc add_diff_inverse)
140apply (simp (no_asm_simp) add: take_succ)
141done
142
143lemma giv_Bounded_lemma2:
144"[| G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel)) |]
145  ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state. s`NbR \<le> length(s`rel)} \<inter>
146   {s\<in>state. s`available_tok #+ tokens(s`giv) =
147    NbT #+ tokens(take(s`NbR, s`rel))})"
148apply (cut_tac giv_Bounded_lamma1)
149apply (cut_tac alloc_prog_preserves_rel_ask_tok)
150apply (auto simp add: Collect_conj_eq [symmetric] alloc_prog_ok_iff)
151apply (subgoal_tac "G \<in> preserves (fun_pair (lift (available_tok), fun_pair (lift (NbR), lift (giv))))")
152apply (rotate_tac -1)
153apply (cut_tac A = "nat * nat * list(nat)"
154             and P = "%<m,n,l> y. n \<le> length(y) &
155                                  m #+ tokens(l) = NbT #+ tokens(take(n,y))"
156             and g = "lift(rel)" and F = alloc_prog
157       in stable_Join_Stable)
158prefer 3 apply assumption
159apply (auto simp add: Collect_conj_eq)
160apply (frule_tac g = length in imp_Increasing_comp)
161apply (blast intro: mono_length)
162apply (auto simp add: refl_prefix)
163apply (drule_tac a=xa and f = "length comp lift(rel)" in Increasing_imp_Stable)
164apply assumption
165apply (auto simp add: Le_def length_type)
166apply (auto dest: ActsD simp add: Stable_def Constrains_def constrains_def)
167apply (drule_tac f = "lift (rel) " in preserves_imp_eq)
168apply assumption+
169apply (force dest: ActsD)
170apply (erule_tac V = "\<forall>x \<in> Acts (alloc_prog) \<union> Acts (G). P(x)" for P in thin_rl)
171apply (erule_tac V = "alloc_prog \<in> stable (u)" for u in thin_rl)
172apply (drule_tac a = "xc`rel" and f = "lift (rel)" in Increasing_imp_Stable)
173apply (auto simp add: Stable_def Constrains_def constrains_def)
174apply (drule bspec, force)
175apply (drule subsetD)
176apply (rule imageI, assumption)
177apply (auto simp add: prefix_take_iff)
178apply (rotate_tac -1)
179apply (erule ssubst)
180apply (auto simp add: take_take min_def)
181done
182
183(*Property (29), page 18:
184  the number of tokens in circulation never exceeds NbT*)
185lemma alloc_prog_giv_Bounded: "alloc_prog \<in> Incr(lift(rel))
186      guarantees Always({s\<in>state. tokens(s`giv) \<le> NbT #+ tokens(s`rel)})"
187apply (cut_tac NbT_pos)
188apply (auto simp add: guar_def)
189apply (rule Always_weaken)
190apply (rule AlwaysI)
191apply (rule_tac [2] giv_Bounded_lemma2, auto)
192apply (rule_tac j = "NbT #+ tokens(take (x` NbR, x`rel))" in le_trans)
193apply (erule subst)
194apply (auto intro!: tokens_mono simp add: prefix_take_iff min_def length_take)
195done
196
197(*Property (30), page 18: the number of tokens given never exceeds the number
198  asked for*)
199lemma alloc_prog_ask_prefix_giv:
200     "alloc_prog \<in> Incr(lift(ask)) guarantees
201                   Always({s\<in>state. <s`giv, s`ask> \<in> prefix(tokbag)})"
202apply (auto intro!: AlwaysI simp add: guar_def)
203apply (subgoal_tac "G \<in> preserves (lift (giv))")
204 prefer 2 apply (simp add: alloc_prog_ok_iff)
205apply (rule_tac P = "%x y. <x,y> \<in> prefix(tokbag)" and A = "list(nat)"
206       in stable_Join_Stable)
207apply safety
208 prefer 2 apply (simp add: lift_def, clarify)
209apply (drule_tac a = k in Increasing_imp_Stable, auto)
210done
211
212subsection\<open>Towards proving the liveness property, (31)\<close>
213
214subsubsection\<open>First, we lead up to a proof of Lemma 49, page 28.\<close>
215
216lemma alloc_prog_transient_lemma:
217     "[|G \<in> program; k\<in>nat|]
218      ==> alloc_prog \<squnion> G \<in>
219             transient({s\<in>state. k \<le> length(s`rel)} \<inter>
220             {s\<in>state. succ(s`NbR) = k})"
221apply auto
222apply (erule_tac V = "G\<notin>u" for u in thin_rl)
223apply (rule_tac act = alloc_rel_act in transientI)
224apply (simp (no_asm) add: alloc_prog_def [THEN def_prg_Acts])
225apply (simp (no_asm) add: alloc_rel_act_def [THEN def_act_eq, THEN act_subset])
226apply (auto simp add: alloc_prog_def [THEN def_prg_Acts] domain_def)
227apply (rule ReplaceI)
228apply (rule_tac x = "x (available_tok:= x`available_tok #+ nth (x`NbR, x`rel),
229                        NbR:=succ (x`NbR))"
230       in exI)
231apply (auto intro!: state_update_type)
232done
233
234lemma alloc_prog_rel_Stable_NbR_lemma:
235    "[| G \<in> program; alloc_prog ok G; k\<in>nat |]
236     ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> succ(s ` NbR)})"
237apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety, auto)
238apply (blast intro: le_trans leI)
239apply (drule_tac f = "lift (NbR)" and A = nat in preserves_imp_increasing)
240apply (drule_tac [2] g = succ in imp_increasing_comp)
241apply (rule_tac [2] mono_succ)
242apply (drule_tac [4] x = k in increasing_imp_stable)
243    prefer 5 apply (simp add: Le_def comp_def, auto)
244done
245
246lemma alloc_prog_NbR_LeadsTo_lemma:
247     "[| G \<in> program; alloc_prog ok G;
248         alloc_prog \<squnion> G \<in> Incr(lift(rel)); k\<in>nat |]
249      ==> alloc_prog \<squnion> G \<in>
250            {s\<in>state. k \<le> length(s`rel)} \<inter> {s\<in>state. succ(s`NbR) = k}
251            \<longmapsto>w {s\<in>state. k \<le> s`NbR}"
252apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k \<le> length (s`rel)})")
253apply (drule_tac [2] a = k and g1 = length in imp_Increasing_comp [THEN Increasing_imp_Stable])
254apply (rule_tac [2] mono_length)
255    prefer 3 apply simp
256apply (simp_all add: refl_prefix Le_def comp_def length_type)
257apply (rule LeadsTo_weaken)
258apply (rule PSP_Stable)
259prefer 2 apply assumption
260apply (rule PSP_Stable)
261apply (rule_tac [2] alloc_prog_rel_Stable_NbR_lemma)
262apply (rule alloc_prog_transient_lemma [THEN transient_imp_leadsTo, THEN leadsTo_imp_LeadsTo], assumption+)
263apply (auto dest: not_lt_imp_le elim: lt_asym simp add: le_iff)
264done
265
266lemma alloc_prog_NbR_LeadsTo_lemma2 [rule_format]:
267    "[| G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
268        k\<in>nat; n \<in> nat; n < k |]
269      ==> alloc_prog \<squnion> G \<in>
270            {s\<in>state . k \<le> length(s ` rel)} \<inter> {s\<in>state . s ` NbR = n}
271               \<longmapsto>w {x \<in> state. k \<le> length(x`rel)} \<inter>
272                 (\<Union>m \<in> greater_than(n). {x \<in> state. x ` NbR=m})"
273apply (unfold greater_than_def)
274apply (rule_tac A' = "{x \<in> state. k \<le> length(x`rel)} \<inter> {x \<in> state. n < x`NbR}"
275       in LeadsTo_weaken_R)
276apply safe
277apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k \<le> length (s`rel) }) ")
278apply (drule_tac [2] a = k and g1 = length in imp_Increasing_comp [THEN Increasing_imp_Stable])
279apply (rule_tac [2] mono_length)
280    prefer 3 apply simp
281apply (simp_all add: refl_prefix Le_def comp_def length_type)
282apply (subst Int_commute [of _ "{x \<in> state . n < x ` NbR}"])
283apply (rule_tac A = "({s \<in> state . k \<le> length (s ` rel) } \<inter>
284                      {s\<in>state . s ` NbR = n}) \<inter> {s\<in>state. k \<le> length(s`rel)}"
285       in LeadsTo_weaken_L)
286apply (rule PSP_Stable, safe)
287apply (rule_tac B = "{x \<in> state . n < length (x ` rel) } \<inter> {s\<in>state . s ` NbR = n}" in LeadsTo_Trans)
288apply (rule_tac [2] LeadsTo_weaken)
289apply (rule_tac [2] k = "succ (n)" in alloc_prog_NbR_LeadsTo_lemma)
290apply simp_all
291apply (rule subset_imp_LeadsTo, auto)
292apply (blast intro: lt_trans2)
293done
294
295lemma Collect_vimage_eq: "u\<in>nat ==> {<s,f(s)>. s \<in> A} -`` u = {s\<in>A. f(s) < u}"
296by (force simp add: lt_def)
297
298(* Lemma 49, page 28 *)
299
300lemma alloc_prog_NbR_LeadsTo_lemma3:
301  "[|G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
302     k\<in>nat|]
303   ==> alloc_prog \<squnion> G \<in>
304           {s\<in>state. k \<le> length(s`rel)} \<longmapsto>w {s\<in>state. k \<le> s`NbR}"
305(* Proof by induction over the difference between k and n *)
306apply (rule_tac f = "\<lambda>s\<in>state. k #- s`NbR" in LessThan_induct)
307apply (simp_all add: lam_def, auto)
308apply (rule single_LeadsTo_I, auto)
309apply (simp (no_asm_simp) add: Collect_vimage_eq)
310apply (rename_tac "s0")
311apply (case_tac "s0`NbR < k")
312apply (rule_tac [2] subset_imp_LeadsTo, safe)
313apply (auto dest!: not_lt_imp_le)
314apply (rule LeadsTo_weaken)
315apply (rule_tac n = "s0`NbR" in alloc_prog_NbR_LeadsTo_lemma2, safe)
316prefer 3 apply assumption
317apply (auto split: nat_diff_split simp add: greater_than_def not_lt_imp_le not_le_iff_lt)
318apply (blast dest: lt_asym)
319apply (force dest: add_lt_elim2)
320done
321
322subsubsection\<open>Towards proving lemma 50, page 29\<close>
323
324lemma alloc_prog_giv_Ensures_lemma:
325"[| G \<in> program; k\<in>nat; alloc_prog ok G;
326  alloc_prog \<squnion> G \<in> Incr(lift(ask)) |] ==>
327  alloc_prog \<squnion> G \<in>
328  {s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
329  {s\<in>state.  k < length(s`ask)} \<inter> {s\<in>state. length(s`giv)=k}
330  Ensures {s\<in>state. ~ k <length(s`ask)} \<union> {s\<in>state. length(s`giv) \<noteq> k}"
331apply (rule EnsuresI, auto)
332apply (erule_tac [2] V = "G\<notin>u" for u in thin_rl)
333apply (rule_tac [2] act = alloc_giv_act in transientI)
334 prefer 2
335 apply (simp add: alloc_prog_def [THEN def_prg_Acts])
336 apply (simp add: alloc_giv_act_def [THEN def_act_eq, THEN act_subset])
337apply (auto simp add: alloc_prog_def [THEN def_prg_Acts] domain_def)
338apply (erule_tac [2] swap)
339apply (rule_tac [2] ReplaceI)
340apply (rule_tac [2] x = "x (giv := x ` giv @ [nth (length(x`giv), x ` ask) ], available_tok := x ` available_tok #- nth (length(x`giv), x ` ask))" in exI)
341apply (auto intro!: state_update_type simp add: app_type)
342apply (rule_tac A = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<inter> {s\<in>state . k < length(s ` ask) } \<inter> {s\<in>state. length(s`giv) =k}" and A' = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<union> {s\<in>state. ~ k < length(s`ask) } \<union> {s\<in>state . length(s ` giv) \<noteq> k}" in Constrains_weaken)
343apply (auto dest: ActsD simp add: Constrains_def constrains_def alloc_prog_def [THEN def_prg_Acts] alloc_prog_ok_iff)
344apply (subgoal_tac "length(xa ` giv @ [nth (length(xa ` giv), xa ` ask) ]) = length(xa ` giv) #+ 1")
345apply (rule_tac [2] trans)
346apply (rule_tac [2] length_app, auto)
347apply (rule_tac j = "xa ` available_tok" in le_trans, auto)
348apply (drule_tac f = "lift (available_tok)" in preserves_imp_eq)
349apply assumption+
350apply auto
351apply (drule_tac a = "xa ` ask" and r = "prefix(tokbag)" and A = "list(tokbag)"
352       in Increasing_imp_Stable)
353apply (auto simp add: prefix_iff)
354apply (drule StableD)
355apply (auto simp add: Constrains_def constrains_def, force)
356done
357
358lemma alloc_prog_giv_Stable_lemma:
359"[| G \<in> program; alloc_prog ok G; k\<in>nat |]
360  ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> length(s`giv)})"
361apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety)
362apply (auto intro: leI)
363apply (drule_tac f = "lift (giv)" and g = length in imp_preserves_comp)
364apply (drule_tac f = "length comp lift (giv)" and A = nat and r = Le in preserves_imp_increasing)
365apply (drule_tac [2] x = k in increasing_imp_stable)
366 prefer 3 apply (simp add: Le_def comp_def)
367apply (auto simp add: length_type)
368done
369
370(* Lemma 50, page 29 *)
371
372lemma alloc_prog_giv_LeadsTo_lemma:
373"[| G \<in> program; alloc_prog ok G;
374    alloc_prog \<squnion> G \<in> Incr(lift(ask)); k\<in>nat |]
375 ==> alloc_prog \<squnion> G \<in>
376        {s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
377        {s\<in>state.  k < length(s`ask)} \<inter>
378        {s\<in>state. length(s`giv) = k}
379        \<longmapsto>w {s\<in>state. k < length(s`giv)}"
380apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} \<longmapsto>w {s\<in>state. ~ k <length(s`ask) } \<union> {s\<in>state. length(s`giv) \<noteq> k}")
381prefer 2 apply (blast intro: alloc_prog_giv_Ensures_lemma [THEN LeadsTo_Basis])
382apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k < length(s`ask) }) ")
383apply (drule PSP_Stable, assumption)
384apply (rule LeadsTo_weaken)
385apply (rule PSP_Stable)
386apply (rule_tac [2] k = k in alloc_prog_giv_Stable_lemma)
387apply (auto simp add: le_iff)
388apply (drule_tac a = "succ (k)" and g1 = length in imp_Increasing_comp [THEN Increasing_imp_Stable])
389apply (rule mono_length)
390 prefer 2 apply simp
391apply (simp_all add: refl_prefix Le_def comp_def length_type)
392done
393
394
395text\<open>Lemma 51, page 29.
396  This theorem states as invariant that if the number of
397  tokens given does not exceed the number returned, then the upper limit
398  (@{term NbT}) does not exceed the number currently available.\<close>
399lemma alloc_prog_Always_lemma:
400"[| G \<in> program; alloc_prog ok G;
401    alloc_prog \<squnion> G \<in> Incr(lift(ask));
402    alloc_prog \<squnion> G \<in> Incr(lift(rel)) |]
403  ==> alloc_prog \<squnion> G \<in>
404        Always({s\<in>state. tokens(s`giv) \<le> tokens(take(s`NbR, s`rel)) \<longrightarrow>
405                NbT \<le> s`available_tok})"
406apply (subgoal_tac
407       "alloc_prog \<squnion> G
408          \<in> Always ({s\<in>state. s`NbR \<le> length(s`rel) } \<inter>
409                    {s\<in>state. s`available_tok #+ tokens(s`giv) = 
410                              NbT #+ tokens(take (s`NbR, s`rel))})")
411apply (rule_tac [2] AlwaysI)
412apply (rule_tac [3] giv_Bounded_lemma2, auto)
413apply (rule Always_weaken, assumption, auto)
414apply (subgoal_tac "0 \<le> tokens(take (x ` NbR, x ` rel)) #- tokens(x`giv) ")
415 prefer 2 apply (force)
416apply (subgoal_tac "x`available_tok =
417                    NbT #+ (tokens(take(x`NbR,x`rel)) #- tokens(x`giv))")
418apply (simp add: )
419apply (auto split: nat_diff_split dest: lt_trans2)
420done
421
422
423
424subsubsection\<open>Main lemmas towards proving property (31)\<close>
425
426lemma LeadsTo_strength_R:
427    "[|  F \<in> C \<longmapsto>w B'; F \<in> A-C \<longmapsto>w B; B'<=B |] ==> F \<in> A \<longmapsto>w  B"
428by (blast intro: LeadsTo_weaken LeadsTo_Un_Un)
429
430lemma PSP_StableI:
431"[| F \<in> Stable(C); F \<in> A - C \<longmapsto>w B;
432   F \<in> A \<inter> C \<longmapsto>w B \<union> (state - C) |] ==> F \<in> A \<longmapsto>w  B"
433apply (rule_tac A = " (A-C) \<union> (A \<inter> C)" in LeadsTo_weaken_L)
434 prefer 2 apply blast
435apply (rule LeadsTo_Un, assumption)
436apply (blast intro: LeadsTo_weaken dest: PSP_Stable)
437done
438
439lemma state_compl_eq [simp]: "state - {s\<in>state. P(s)} = {s\<in>state. ~P(s)}"
440by auto
441
442(*needed?*)
443lemma single_state_Diff_eq [simp]: "{s}-{x \<in> state. P(x)} = (if s\<in>state & P(s) then 0 else {s})"
444by auto
445
446
447locale alloc_progress =
448 fixes G
449 assumes Gprog [intro,simp]: "G \<in> program"
450     and okG [iff]:          "alloc_prog ok G"
451     and Incr_rel [intro]:   "alloc_prog \<squnion> G \<in> Incr(lift(rel))"
452     and Incr_ask [intro]:   "alloc_prog \<squnion> G \<in> Incr(lift(ask))"
453     and safety:   "alloc_prog \<squnion> G
454                      \<in> Always(\<Inter>k \<in> nat. {s\<in>state. nth(k, s`ask) \<le> NbT})"
455     and progress: "alloc_prog \<squnion> G
456                      \<in> (\<Inter>k\<in>nat. {s\<in>state. k \<le> tokens(s`giv)} \<longmapsto>w
457                        {s\<in>state. k \<le> tokens(s`rel)})"
458
459(*First step in proof of (31) -- the corrected version from Charpentier.
460  This lemma implies that if a client releases some tokens then the Allocator
461  will eventually recognize that they've been released.*)
462lemma (in alloc_progress) tokens_take_NbR_lemma:
463 "k \<in> tokbag
464  ==> alloc_prog \<squnion> G \<in>
465        {s\<in>state. k \<le> tokens(s`rel)}
466        \<longmapsto>w {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
467apply (rule single_LeadsTo_I, safe)
468apply (rule_tac a1 = "s`rel" in Increasing_imp_Stable [THEN PSP_StableI])
469apply (rule_tac [4] k1 = "length(s`rel)" in alloc_prog_NbR_LeadsTo_lemma3 [THEN LeadsTo_strength_R])
470apply (rule_tac [8] subset_imp_LeadsTo)
471apply (auto intro!: Incr_rel)
472apply (rule_tac j = "tokens(take (length(s`rel), x`rel))" in le_trans)
473apply (rule_tac j = "tokens(take (length(s`rel), s`rel))" in le_trans)
474apply (auto intro!: tokens_mono take_mono simp add: prefix_iff)
475done
476
477(*** Rest of proofs done by lcp ***)
478
479(*Second step in proof of (31): by LHS of the guarantee and transivity of
480  \<longmapsto>w *)
481lemma (in alloc_progress) tokens_take_NbR_lemma2:
482     "k \<in> tokbag
483      ==> alloc_prog \<squnion> G \<in>
484            {s\<in>state. tokens(s`giv) = k}
485            \<longmapsto>w {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
486apply (rule LeadsTo_Trans)
487 apply (rule_tac [2] tokens_take_NbR_lemma)
488 prefer 2 apply assumption
489apply (insert progress) 
490apply (blast intro: LeadsTo_weaken_L progress nat_into_Ord)
491done
492
493(*Third step in proof of (31): by PSP with the fact that giv increases *)
494lemma (in alloc_progress) length_giv_disj:
495     "[| k \<in> tokbag; n \<in> nat |]
496      ==> alloc_prog \<squnion> G \<in>
497            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
498            \<longmapsto>w
499              {s\<in>state. (length(s`giv) = n & tokens(s`giv) = k &
500                         k \<le> tokens(take(s`NbR, s`rel))) | n < length(s`giv)}"
501apply (rule single_LeadsTo_I, safe)
502apply (rule_tac a1 = "s`giv" in Increasing_imp_Stable [THEN PSP_StableI])
503apply (rule alloc_prog_Increasing_giv [THEN guaranteesD])
504apply (simp_all add: Int_cons_left)
505apply (rule LeadsTo_weaken)
506apply (rule_tac k = "tokens(s`giv)" in tokens_take_NbR_lemma2)
507apply auto
508apply (force dest: prefix_length_le [THEN le_iff [THEN iffD1]]) 
509apply (simp add: not_lt_iff_le)
510apply (force dest: prefix_length_le_equal) 
511done
512
513(*Fourth step in proof of (31): we apply lemma (51) *)
514lemma (in alloc_progress) length_giv_disj2:
515     "[|k \<in> tokbag; n \<in> nat|]
516      ==> alloc_prog \<squnion> G \<in>
517            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
518            \<longmapsto>w
519              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
520                        n < length(s`giv)}"
521apply (rule LeadsTo_weaken_R)
522apply (rule Always_LeadsToD [OF alloc_prog_Always_lemma length_giv_disj], auto)
523done
524
525(*Fifth step in proof of (31): from the fourth step, taking the union over all
526  k\<in>nat *)
527lemma (in alloc_progress) length_giv_disj3:
528     "n \<in> nat
529      ==> alloc_prog \<squnion> G \<in>
530            {s\<in>state. length(s`giv) = n}
531            \<longmapsto>w
532              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
533                        n < length(s`giv)}"
534apply (rule LeadsTo_weaken_L)
535apply (rule_tac I = nat in LeadsTo_UN)
536apply (rule_tac k = i in length_giv_disj2)
537apply (simp_all add: UN_conj_eq)
538done
539
540(*Sixth step in proof of (31): from the fifth step, by PSP with the
541  assumption that ask increases *)
542lemma (in alloc_progress) length_ask_giv:
543 "[|k \<in> nat;  n < k|]
544  ==> alloc_prog \<squnion> G \<in>
545        {s\<in>state. length(s`ask) = k & length(s`giv) = n}
546        \<longmapsto>w
547          {s\<in>state. (NbT \<le> s`available_tok & length(s`giv) < length(s`ask) &
548                     length(s`giv) = n) |
549                    n < length(s`giv)}"
550apply (rule single_LeadsTo_I, safe)
551apply (rule_tac a1 = "s`ask" and f1 = "lift(ask)" 
552       in Increasing_imp_Stable [THEN PSP_StableI])
553apply (rule Incr_ask, simp_all)
554apply (rule LeadsTo_weaken)
555apply (rule_tac n = "length(s ` giv)" in length_giv_disj3)
556apply simp_all
557apply blast
558apply clarify
559apply simp
560apply (blast dest!: prefix_length_le intro: lt_trans2)
561done
562
563
564(*Seventh step in proof of (31): no request (ask[k]) exceeds NbT *)
565lemma (in alloc_progress) length_ask_giv2:
566     "[|k \<in> nat;  n < k|]
567      ==> alloc_prog \<squnion> G \<in>
568            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
569            \<longmapsto>w
570              {s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok &
571                         length(s`giv) < length(s`ask) & length(s`giv) = n) |
572                        n < length(s`giv)}"
573apply (rule LeadsTo_weaken_R)
574apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify)
575apply (simp add: INT_iff)
576apply (drule_tac x = "length(x ` giv)" and P = "%x. f (x) \<le> NbT" for f in bspec)
577apply simp
578apply (blast intro: le_trans)
579done
580
581(*Eighth step in proof of (31): by 50, we get |giv| > n. *)
582lemma (in alloc_progress) extend_giv:
583     "[| k \<in> nat;  n < k|]
584      ==> alloc_prog \<squnion> G \<in>
585            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
586            \<longmapsto>w {s\<in>state. n < length(s`giv)}"
587apply (rule LeadsTo_Un_duplicate)
588apply (rule LeadsTo_cancel1)
589apply (rule_tac [2] alloc_prog_giv_LeadsTo_lemma)
590apply (simp_all add: Incr_ask lt_nat_in_nat)
591apply (rule LeadsTo_weaken_R)
592apply (rule length_ask_giv2, auto)
593done
594
595(*Ninth and tenth steps in proof of (31): by 50, we get |giv| > n.
596  The report has an error: putting |ask|=k for the precondition fails because
597  we can't expect |ask| to remain fixed until |giv| increases.*)
598lemma (in alloc_progress) alloc_prog_ask_LeadsTo_giv:
599 "k \<in> nat
600  ==> alloc_prog \<squnion> G \<in>
601        {s\<in>state. k \<le> length(s`ask)} \<longmapsto>w {s\<in>state. k \<le> length(s`giv)}"
602(* Proof by induction over the difference between k and n *)
603apply (rule_tac f = "\<lambda>s\<in>state. k #- length(s`giv)" in LessThan_induct)
604apply (auto simp add: lam_def Collect_vimage_eq)
605apply (rule single_LeadsTo_I, auto)
606apply (rename_tac "s0")
607apply (case_tac "length(s0 ` giv) < length(s0 ` ask) ")
608 apply (rule_tac [2] subset_imp_LeadsTo)
609  apply (auto simp add: not_lt_iff_le)
610 prefer 2 apply (blast dest: le_imp_not_lt intro: lt_trans2)
611apply (rule_tac a1 = "s0`ask" and f1 = "lift (ask)"
612       in Increasing_imp_Stable [THEN PSP_StableI])
613apply (rule Incr_ask, simp)
614apply (force)
615apply (rule LeadsTo_weaken)
616apply (rule_tac n = "length(s0 ` giv)" and k = "length(s0 ` ask)"
617       in extend_giv) 
618apply (auto dest: not_lt_imp_le simp add: leI diff_lt_iff_lt) 
619apply (blast dest!: prefix_length_le intro: lt_trans2)
620done
621
622(*Final lemma: combine previous result with lemma (30)*)
623lemma (in alloc_progress) final:
624     "h \<in> list(tokbag)
625      ==> alloc_prog \<squnion> G
626            \<in> {s\<in>state. <h, s`ask> \<in> prefix(tokbag)} \<longmapsto>w
627              {s\<in>state. <h, s`giv> \<in> prefix(tokbag)}"
628apply (rule single_LeadsTo_I)
629 prefer 2 apply simp
630apply (rename_tac s0)
631apply (rule_tac a1 = "s0`ask" and f1 = "lift (ask)"
632       in Increasing_imp_Stable [THEN PSP_StableI])
633   apply (rule Incr_ask)
634  apply (simp_all add: Int_cons_left)
635apply (rule LeadsTo_weaken)
636apply (rule_tac k1 = "length(s0 ` ask)"
637       in Always_LeadsToD [OF alloc_prog_ask_prefix_giv [THEN guaranteesD]
638                              alloc_prog_ask_LeadsTo_giv])
639apply (auto simp add: Incr_ask)
640apply (blast intro: length_le_prefix_imp_prefix prefix_trans prefix_length_le 
641                    lt_trans2)
642done
643
644(** alloc_prog liveness property (31), page 18 **)
645
646theorem alloc_prog_progress:
647"alloc_prog \<in>
648    Incr(lift(ask)) \<inter> Incr(lift(rel)) \<inter>
649    Always(\<Inter>k \<in> nat. {s\<in>state. nth(k, s`ask) \<le> NbT}) \<inter>
650    (\<Inter>k\<in>nat. {s\<in>state. k \<le> tokens(s`giv)} \<longmapsto>w
651              {s\<in>state. k \<le> tokens(s`rel)})
652  guarantees (\<Inter>h \<in> list(tokbag).
653              {s\<in>state. <h, s`ask> \<in> prefix(tokbag)} \<longmapsto>w
654              {s\<in>state. <h, s`giv> \<in> prefix(tokbag)})"
655apply (rule guaranteesI)
656apply (rule INT_I)
657apply (rule alloc_progress.final)
658apply (auto simp add: alloc_progress_def)
659done
660
661end
662