author | wenzelm |
Sat, 28 Jun 2008 21:21:13 +0200 | |
changeset 27381 | 19ae7064f00f |
parent 24893 | b8ef7afe3a6b |
child 32960 | 69916a850301 |
permissions | -rw-r--r-- |
14052 | 1 |
(* Title: ZF/UNITY/AllocBase.thy |
2 |
ID: $Id$ |
|
3 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory |
|
4 |
Copyright 2001 University of Cambridge |
|
5 |
||
6 |
*) |
|
7 |
||
14076 | 8 |
header{*Common declarations for Chandy and Charpentier's Allocator*} |
9 |
||
16417 | 10 |
theory AllocBase imports Follows MultisetSum Guar begin |
14052 | 11 |
|
12 |
consts |
|
13 |
NbT :: i (* Number of tokens in system *) |
|
14 |
Nclients :: i (* Number of clients *) |
|
14076 | 15 |
|
24892 | 16 |
abbreviation (input) |
17 |
tokbag :: i (* tokbags could be multisets...or any ordered type?*) |
|
18 |
where |
|
19 |
"tokbag == nat" |
|
14076 | 20 |
|
21 |
axioms |
|
22 |
NbT_pos: "NbT \<in> nat-{0}" |
|
23 |
Nclients_pos: "Nclients \<in> nat-{0}" |
|
14052 | 24 |
|
14076 | 25 |
text{*This function merely sums the elements of a list*} |
26 |
consts tokens :: "i =>i" |
|
14052 | 27 |
item :: i (* Items to be merged/distributed *) |
28 |
primrec |
|
29 |
"tokens(Nil) = 0" |
|
30 |
"tokens (Cons(x,xs)) = x #+ tokens(xs)" |
|
31 |
||
32 |
consts |
|
14076 | 33 |
bag_of :: "i => i" |
14052 | 34 |
|
35 |
primrec |
|
36 |
"bag_of(Nil) = 0" |
|
37 |
"bag_of(Cons(x,xs)) = {#x#} +# bag_of(xs)" |
|
38 |
||
14076 | 39 |
|
40 |
text{*Definitions needed in Client.thy. We define a recursive predicate |
|
41 |
using 0 and 1 to code the truth values.*} |
|
42 |
consts all_distinct0 :: "i=>i" |
|
43 |
primrec |
|
14052 | 44 |
"all_distinct0(Nil) = 1" |
45 |
"all_distinct0(Cons(a, l)) = |
|
14076 | 46 |
(if a \<in> set_of_list(l) then 0 else all_distinct0(l))" |
14052 | 47 |
|
24893 | 48 |
definition |
49 |
all_distinct :: "i=>o" where |
|
14076 | 50 |
"all_distinct(l) == all_distinct0(l)=1" |
14052 | 51 |
|
24893 | 52 |
definition |
53 |
state_of :: "i =>i" --{* coersion from anyting to state *} where |
|
14076 | 54 |
"state_of(s) == if s \<in> state then s else st0" |
14052 | 55 |
|
24893 | 56 |
definition |
57 |
lift :: "i =>(i=>i)" --{* simplifies the expression of programs*} where |
|
14076 | 58 |
"lift(x) == %s. s`x" |
14052 | 59 |
|
14076 | 60 |
text{* function to show that the set of variables is infinite *} |
61 |
consts |
|
62 |
nat_list_inj :: "i=>i" |
|
63 |
var_inj :: "i=>i" |
|
14052 | 64 |
|
65 |
primrec |
|
66 |
"nat_list_inj(0) = Nil" |
|
67 |
"nat_list_inj(succ(n)) = Cons(n, nat_list_inj(n))" |
|
68 |
||
69 |
primrec |
|
70 |
"var_inj(Var(l)) = length(l)" |
|
71 |
||
24893 | 72 |
definition |
73 |
nat_var_inj :: "i=>i" where |
|
74 |
"nat_var_inj(n) == Var(nat_list_inj(n))" |
|
75 |
||
14076 | 76 |
|
77 |
subsection{*Various simple lemmas*} |
|
78 |
||
79 |
lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients & 0 < NbT" |
|
80 |
apply (cut_tac Nclients_pos NbT_pos) |
|
81 |
apply (auto intro: Ord_0_lt) |
|
82 |
done |
|
83 |
||
84 |
lemma Nclients_NbT_not_0 [simp]: "Nclients \<noteq> 0 & NbT \<noteq> 0" |
|
85 |
by (cut_tac Nclients_pos NbT_pos, auto) |
|
86 |
||
87 |
lemma Nclients_type [simp,TC]: "Nclients\<in>nat" |
|
88 |
by (cut_tac Nclients_pos NbT_pos, auto) |
|
89 |
||
90 |
lemma NbT_type [simp,TC]: "NbT\<in>nat" |
|
91 |
by (cut_tac Nclients_pos NbT_pos, auto) |
|
92 |
||
93 |
lemma INT_Nclient_iff [iff]: |
|
94 |
"b\<in>Inter(RepFun(Nclients, B)) <-> (\<forall>x\<in>Nclients. b\<in>B(x))" |
|
14095
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents:
14084
diff
changeset
|
95 |
by (force simp add: INT_iff) |
14076 | 96 |
|
97 |
lemma setsum_fun_mono [rule_format]: |
|
98 |
"n\<in>nat ==> |
|
99 |
(\<forall>i\<in>nat. i<n --> f(i) $<= g(i)) --> |
|
100 |
setsum(f, n) $<= setsum(g,n)" |
|
101 |
apply (induct_tac "n", simp_all) |
|
102 |
apply (subgoal_tac "Finite(x) & x\<notin>x") |
|
103 |
prefer 2 apply (simp add: nat_into_Finite mem_not_refl, clarify) |
|
104 |
apply (simp (no_asm_simp) add: succ_def) |
|
105 |
apply (subgoal_tac "\<forall>i\<in>nat. i<x--> f(i) $<= g(i) ") |
|
106 |
prefer 2 apply (force dest: leI) |
|
107 |
apply (rule zadd_zle_mono, simp_all) |
|
108 |
done |
|
109 |
||
110 |
lemma tokens_type [simp,TC]: "l\<in>list(A) ==> tokens(l)\<in>nat" |
|
111 |
by (erule list.induct, auto) |
|
112 |
||
113 |
lemma tokens_mono_aux [rule_format]: |
|
114 |
"xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A) |
|
115 |
--> tokens(xs) \<le> tokens(ys)" |
|
116 |
apply (induct_tac "xs") |
|
117 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def) |
|
118 |
done |
|
119 |
||
120 |
lemma tokens_mono: "<xs, ys>\<in>prefix(A) ==> tokens(xs) \<le> tokens(ys)" |
|
121 |
apply (cut_tac prefix_type) |
|
122 |
apply (blast intro: tokens_mono_aux) |
|
123 |
done |
|
124 |
||
125 |
lemma mono_tokens [iff]: "mono1(list(A), prefix(A), nat, Le,tokens)" |
|
126 |
apply (unfold mono1_def) |
|
127 |
apply (auto intro: tokens_mono simp add: Le_def) |
|
128 |
done |
|
129 |
||
130 |
lemma tokens_append [simp]: |
|
131 |
"[| xs\<in>list(A); ys\<in>list(A) |] ==> tokens(xs@ys) = tokens(xs) #+ tokens(ys)" |
|
132 |
apply (induct_tac "xs", auto) |
|
133 |
done |
|
134 |
||
135 |
subsection{*The function @{term bag_of}*} |
|
136 |
||
137 |
lemma bag_of_type [simp,TC]: "l\<in>list(A) ==>bag_of(l)\<in>Mult(A)" |
|
138 |
apply (induct_tac "l") |
|
139 |
apply (auto simp add: Mult_iff_multiset) |
|
140 |
done |
|
141 |
||
142 |
lemma bag_of_multiset: |
|
143 |
"l\<in>list(A) ==> multiset(bag_of(l)) & mset_of(bag_of(l))<=A" |
|
144 |
apply (drule bag_of_type) |
|
145 |
apply (auto simp add: Mult_iff_multiset) |
|
146 |
done |
|
147 |
||
148 |
lemma bag_of_append [simp]: |
|
149 |
"[| xs\<in>list(A); ys\<in>list(A)|] ==> bag_of(xs@ys) = bag_of(xs) +# bag_of(ys)" |
|
150 |
apply (induct_tac "xs") |
|
151 |
apply (auto simp add: bag_of_multiset munion_assoc) |
|
152 |
done |
|
153 |
||
154 |
lemma bag_of_mono_aux [rule_format]: |
|
155 |
"xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A) |
|
156 |
--> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)" |
|
157 |
apply (induct_tac "xs", simp_all, clarify) |
|
158 |
apply (frule_tac l = ys in bag_of_multiset) |
|
159 |
apply (auto intro: empty_le_MultLe simp add: prefix_def) |
|
160 |
apply (rule munion_mono) |
|
161 |
apply (force simp add: MultLe_def Mult_iff_multiset) |
|
162 |
apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) |
|
163 |
done |
|
164 |
||
165 |
lemma bag_of_mono [intro]: |
|
166 |
"[| <xs, ys>\<in>prefix(A); xs\<in>list(A); ys\<in>list(A) |] |
|
167 |
==> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)" |
|
168 |
apply (blast intro: bag_of_mono_aux) |
|
169 |
done |
|
170 |
||
171 |
lemma mono_bag_of [simp]: |
|
172 |
"mono1(list(A), prefix(A), Mult(A), MultLe(A,r), bag_of)" |
|
173 |
by (auto simp add: mono1_def bag_of_type) |
|
174 |
||
175 |
||
176 |
subsection{*The function @{term msetsum}*} |
|
177 |
||
178 |
lemmas nat_into_Fin = eqpoll_refl [THEN [2] Fin_lemma] |
|
179 |
||
180 |
lemma list_Int_length_Fin: "l \<in> list(A) ==> C Int length(l) \<in> Fin(length(l))" |
|
181 |
apply (drule length_type) |
|
182 |
apply (rule Fin_subset) |
|
183 |
apply (rule Int_lower2) |
|
184 |
apply (erule nat_into_Fin) |
|
185 |
done |
|
186 |
||
187 |
||
188 |
||
189 |
lemma mem_Int_imp_lt_length: |
|
190 |
"[|xs \<in> list(A); k \<in> C \<inter> length(xs)|] ==> k < length(xs)" |
|
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14076
diff
changeset
|
191 |
by (simp add: ltI) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14076
diff
changeset
|
192 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14076
diff
changeset
|
193 |
lemma Int_succ_right: |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14076
diff
changeset
|
194 |
"A Int succ(k) = (if k : A then cons(k, A Int k) else A Int k)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14076
diff
changeset
|
195 |
by auto |
14076 | 196 |
|
197 |
||
198 |
lemma bag_of_sublist_lemma: |
|
199 |
"[|C \<subseteq> nat; x \<in> A; xs \<in> list(A)|] |
|
200 |
==> msetsum(\<lambda>i. {#nth(i, xs @ [x])#}, C \<inter> succ(length(xs)), A) = |
|
201 |
(if length(xs) \<in> C then |
|
202 |
{#x#} +# msetsum(\<lambda>x. {#nth(x, xs)#}, C \<inter> length(xs), A) |
|
203 |
else msetsum(\<lambda>x. {#nth(x, xs)#}, C \<inter> length(xs), A))" |
|
204 |
apply (simp add: subsetD nth_append lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong) |
|
205 |
apply (simp add: Int_succ_right) |
|
206 |
apply (simp add: lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong, clarify) |
|
207 |
apply (subst msetsum_cons) |
|
208 |
apply (rule_tac [3] succI1) |
|
209 |
apply (blast intro: list_Int_length_Fin subset_succI [THEN Fin_mono, THEN subsetD]) |
|
210 |
apply (simp add: mem_not_refl) |
|
211 |
apply (simp add: nth_type lt_not_refl) |
|
212 |
apply (blast intro: nat_into_Ord ltI length_type) |
|
213 |
apply (simp add: lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong) |
|
214 |
done |
|
215 |
||
216 |
lemma bag_of_sublist_lemma2: |
|
217 |
"l\<in>list(A) ==> |
|
218 |
C <= nat ==> |
|
219 |
bag_of(sublist(l, C)) = |
|
220 |
msetsum(%i. {#nth(i, l)#}, C Int length(l), A)" |
|
221 |
apply (erule list_append_induct) |
|
222 |
apply (simp (no_asm)) |
|
223 |
apply (simp (no_asm_simp) add: sublist_append nth_append bag_of_sublist_lemma munion_commute bag_of_sublist_lemma msetsum_multiset munion_0) |
|
224 |
done |
|
225 |
||
226 |
||
227 |
lemma nat_Int_length_eq: "l \<in> list(A) ==> nat \<inter> length(l) = length(l)" |
|
228 |
apply (rule Int_absorb1) |
|
229 |
apply (rule OrdmemD, auto) |
|
230 |
done |
|
231 |
||
232 |
(*eliminating the assumption C<=nat*) |
|
233 |
lemma bag_of_sublist: |
|
234 |
"l\<in>list(A) ==> |
|
235 |
bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C Int length(l), A)" |
|
236 |
apply (subgoal_tac " bag_of (sublist (l, C Int nat)) = msetsum (%i. {#nth (i, l) #}, C Int length (l), A) ") |
|
237 |
apply (simp add: sublist_Int_eq) |
|
238 |
apply (simp add: bag_of_sublist_lemma2 Int_lower2 Int_assoc nat_Int_length_eq) |
|
239 |
done |
|
240 |
||
241 |
lemma bag_of_sublist_Un_Int: |
|
242 |
"l\<in>list(A) ==> |
|
243 |
bag_of(sublist(l, B Un C)) +# bag_of(sublist(l, B Int C)) = |
|
244 |
bag_of(sublist(l, B)) +# bag_of(sublist(l, C))" |
|
245 |
apply (subgoal_tac "B Int C Int length (l) = (B Int length (l)) Int (C Int length (l))") |
|
246 |
prefer 2 apply blast |
|
247 |
apply (simp (no_asm_simp) add: bag_of_sublist Int_Un_distrib2 msetsum_Un_Int) |
|
248 |
apply (rule msetsum_Un_Int) |
|
249 |
apply (erule list_Int_length_Fin)+ |
|
250 |
apply (simp add: ltI nth_type) |
|
251 |
done |
|
252 |
||
253 |
||
254 |
lemma bag_of_sublist_Un_disjoint: |
|
255 |
"[| l\<in>list(A); B Int C = 0 |] |
|
256 |
==> bag_of(sublist(l, B Un C)) = |
|
257 |
bag_of(sublist(l, B)) +# bag_of(sublist(l, C))" |
|
258 |
by (simp add: bag_of_sublist_Un_Int [symmetric] bag_of_multiset) |
|
259 |
||
260 |
||
261 |
lemma bag_of_sublist_UN_disjoint [rule_format]: |
|
262 |
"[|Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j --> A(i) \<inter> A(j) = 0; |
|
263 |
l\<in>list(B) |] |
|
264 |
==> bag_of(sublist(l, \<Union>i\<in>I. A(i))) = |
|
265 |
(msetsum(%i. bag_of(sublist(l, A(i))), I, B)) " |
|
266 |
apply (simp (no_asm_simp) del: UN_simps |
|
267 |
add: UN_simps [symmetric] bag_of_sublist) |
|
268 |
apply (subst msetsum_UN_disjoint [of _ _ _ "length (l)"]) |
|
269 |
apply (drule Finite_into_Fin, assumption) |
|
270 |
prefer 3 apply force |
|
271 |
apply (auto intro!: Fin_IntI2 Finite_into_Fin simp add: ltI nth_type length_type nat_into_Finite) |
|
272 |
done |
|
273 |
||
274 |
||
275 |
lemma part_ord_Lt [simp]: "part_ord(nat, Lt)" |
|
276 |
apply (unfold part_ord_def Lt_def irrefl_def trans_on_def) |
|
277 |
apply (auto intro: lt_trans) |
|
278 |
done |
|
279 |
||
280 |
subsubsection{*The function @{term all_distinct}*} |
|
281 |
||
282 |
lemma all_distinct_Nil [simp]: "all_distinct(Nil)" |
|
283 |
by (unfold all_distinct_def, auto) |
|
284 |
||
285 |
lemma all_distinct_Cons [simp]: |
|
286 |
"all_distinct(Cons(a,l)) <-> |
|
287 |
(a\<in>set_of_list(l) --> False) & (a \<notin> set_of_list(l) --> all_distinct(l))" |
|
288 |
apply (unfold all_distinct_def) |
|
289 |
apply (auto elim: list.cases) |
|
290 |
done |
|
291 |
||
292 |
subsubsection{*The function @{term state_of}*} |
|
293 |
||
294 |
lemma state_of_state: "s\<in>state ==> state_of(s)=s" |
|
295 |
by (unfold state_of_def, auto) |
|
296 |
declare state_of_state [simp] |
|
297 |
||
298 |
||
299 |
lemma state_of_idem: "state_of(state_of(s))=state_of(s)" |
|
300 |
||
301 |
apply (unfold state_of_def, auto) |
|
302 |
done |
|
303 |
declare state_of_idem [simp] |
|
304 |
||
305 |
||
306 |
lemma state_of_type [simp,TC]: "state_of(s)\<in>state" |
|
307 |
by (unfold state_of_def, auto) |
|
308 |
||
309 |
lemma lift_apply [simp]: "lift(x, s)=s`x" |
|
310 |
by (simp add: lift_def) |
|
311 |
||
312 |
||
313 |
(** Used in ClientImp **) |
|
314 |
||
315 |
lemma gen_Increains_state_of_eq: |
|
316 |
"Increasing(A, r, %s. f(state_of(s))) = Increasing(A, r, f)" |
|
317 |
apply (unfold Increasing_def, auto) |
|
318 |
done |
|
319 |
||
320 |
lemmas Increasing_state_ofD1 = |
|
321 |
gen_Increains_state_of_eq [THEN equalityD1, THEN subsetD, standard] |
|
322 |
lemmas Increasing_state_ofD2 = |
|
323 |
gen_Increains_state_of_eq [THEN equalityD2, THEN subsetD, standard] |
|
324 |
||
325 |
lemma Follows_state_of_eq: |
|
326 |
"Follows(A, r, %s. f(state_of(s)), %s. g(state_of(s))) = |
|
327 |
Follows(A, r, f, g)" |
|
328 |
apply (unfold Follows_def Increasing_def, auto) |
|
329 |
done |
|
330 |
||
331 |
lemmas Follows_state_ofD1 = |
|
332 |
Follows_state_of_eq [THEN equalityD1, THEN subsetD, standard] |
|
333 |
lemmas Follows_state_ofD2 = |
|
334 |
Follows_state_of_eq [THEN equalityD2, THEN subsetD, standard] |
|
335 |
||
336 |
lemma nat_list_inj_type: "n\<in>nat ==> nat_list_inj(n)\<in>list(nat)" |
|
337 |
by (induct_tac "n", auto) |
|
338 |
||
339 |
lemma length_nat_list_inj: "n\<in>nat ==> length(nat_list_inj(n)) = n" |
|
340 |
by (induct_tac "n", auto) |
|
341 |
||
342 |
lemma var_infinite_lemma: |
|
343 |
"(\<lambda>x\<in>nat. nat_var_inj(x))\<in>inj(nat, var)" |
|
344 |
apply (unfold nat_var_inj_def) |
|
345 |
apply (rule_tac d = var_inj in lam_injective) |
|
346 |
apply (auto simp add: var.intros nat_list_inj_type) |
|
347 |
apply (simp add: length_nat_list_inj) |
|
348 |
done |
|
349 |
||
350 |
lemma nat_lepoll_var: "nat lepoll var" |
|
351 |
apply (unfold lepoll_def) |
|
352 |
apply (rule_tac x = " (\<lambda>x\<in>nat. nat_var_inj (x))" in exI) |
|
353 |
apply (rule var_infinite_lemma) |
|
354 |
done |
|
355 |
||
356 |
lemma var_not_Finite: "~Finite(var)" |
|
357 |
apply (insert nat_not_Finite) |
|
358 |
apply (blast intro: lepoll_Finite [OF nat_lepoll_var]) |
|
359 |
done |
|
360 |
||
361 |
lemma not_Finite_imp_exist: "~Finite(A) ==> \<exists>x. x\<in>A" |
|
362 |
apply (subgoal_tac "A\<noteq>0") |
|
363 |
apply (auto simp add: Finite_0) |
|
364 |
done |
|
365 |
||
366 |
lemma Inter_Diff_var_iff: |
|
367 |
"Finite(A) ==> b\<in>(Inter(RepFun(var-A, B))) <-> (\<forall>x\<in>var-A. b\<in>B(x))" |
|
368 |
apply (subgoal_tac "\<exists>x. x\<in>var-A", auto) |
|
369 |
apply (subgoal_tac "~Finite (var-A) ") |
|
370 |
apply (drule not_Finite_imp_exist, auto) |
|
371 |
apply (cut_tac var_not_Finite) |
|
372 |
apply (erule swap) |
|
373 |
apply (rule_tac B = A in Diff_Finite, auto) |
|
374 |
done |
|
375 |
||
376 |
lemma Inter_var_DiffD: |
|
377 |
"[| b\<in>Inter(RepFun(var-A, B)); Finite(A); x\<in>var-A |] ==> b\<in>B(x)" |
|
378 |
by (simp add: Inter_Diff_var_iff) |
|
379 |
||
380 |
(* [| Finite(A); (\<forall>x\<in>var-A. b\<in>B(x)) |] ==> b\<in>Inter(RepFun(var-A, B)) *) |
|
381 |
lemmas Inter_var_DiffI = Inter_Diff_var_iff [THEN iffD2, standard] |
|
382 |
||
383 |
declare Inter_var_DiffI [intro!] |
|
384 |
||
385 |
lemma Acts_subset_Int_Pow_simp [simp]: |
|
386 |
"Acts(F)<= A \<inter> Pow(state*state) <-> Acts(F)<=A" |
|
387 |
by (insert Acts_type [of F], auto) |
|
388 |
||
389 |
lemma setsum_nsetsum_eq: |
|
390 |
"[| Finite(A); \<forall>x\<in>A. g(x)\<in>nat |] |
|
391 |
==> setsum(%x. $#(g(x)), A) = $# nsetsum(%x. g(x), A)" |
|
392 |
apply (erule Finite_induct) |
|
393 |
apply (auto simp add: int_of_add) |
|
394 |
done |
|
395 |
||
396 |
lemma nsetsum_cong: |
|
397 |
"[| A=B; \<forall>x\<in>A. f(x)=g(x); \<forall>x\<in>A. g(x)\<in>nat; Finite(A) |] |
|
398 |
==> nsetsum(f, A) = nsetsum(g, B)" |
|
399 |
apply (subgoal_tac "$# nsetsum (f, A) = $# nsetsum (g, B)", simp) |
|
400 |
apply (simp add: setsum_nsetsum_eq [symmetric] cong: setsum_cong) |
|
401 |
done |
|
402 |
||
14052 | 403 |
end |