| author | blanchet | 
| Tue, 15 Nov 2011 22:13:39 +0100 | |
| changeset 45510 | 96696c360b3e | 
| parent 39246 | 9e58f0499f57 | 
| child 45827 | 66c68453455c | 
| permissions | -rw-r--r-- | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
24147 
diff
changeset
 | 
1  | 
(* Title: HOL/UNITY/Comp/AllocBase.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  | 
|
| 13798 | 6  | 
header{*Common Declarations for Chandy and Charpentier's Allocator*}
 | 
7  | 
||
| 18556 | 8  | 
theory AllocBase imports "../UNITY_Main" begin  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
10  | 
consts  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
11  | 
NbT :: nat (*Number of tokens in system*)  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
12  | 
Nclients :: nat (*Number of clients*)  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
13  | 
|
| 13798 | 14  | 
axioms  | 
15  | 
NbT_pos: "0 < NbT"  | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
17  | 
(*This function merely sums the elements of a list*)  | 
| 39246 | 18  | 
primrec tokens :: "nat list => nat" where  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
19  | 
"tokens [] = 0"  | 
| 39246 | 20  | 
| "tokens (x#xs) = x + tokens xs"  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 39246 | 22  | 
abbreviation (input) "bag_of \<equiv> multiset_of"  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
23  | 
|
| 13798 | 24  | 
lemma setsum_fun_mono [rule_format]:  | 
25  | 
"!!f :: nat=>nat.  | 
|
26  | 
(ALL i. i<n --> f i <= g i) -->  | 
|
27  | 
setsum f (lessThan n) <= setsum g (lessThan n)"  | 
|
28  | 
apply (induct_tac "n")  | 
|
29  | 
apply (auto simp add: lessThan_Suc)  | 
|
30  | 
done  | 
|
31  | 
||
32  | 
lemma tokens_mono_prefix [rule_format]:  | 
|
33  | 
"ALL xs. xs <= ys --> tokens xs <= tokens ys"  | 
|
34  | 
apply (induct_tac "ys")  | 
|
35  | 
apply (auto simp add: prefix_Cons)  | 
|
36  | 
done  | 
|
37  | 
||
38  | 
lemma mono_tokens: "mono tokens"  | 
|
39  | 
apply (unfold mono_def)  | 
|
40  | 
apply (blast intro: tokens_mono_prefix)  | 
|
41  | 
done  | 
|
42  | 
||
43  | 
||
44  | 
(** bag_of **)  | 
|
45  | 
||
46  | 
lemma bag_of_append [simp]: "bag_of (l@l') = bag_of l + bag_of l'"  | 
|
47  | 
apply (induct_tac "l", simp)  | 
|
| 14738 | 48  | 
apply (simp add: add_ac)  | 
| 13798 | 49  | 
done  | 
50  | 
||
51  | 
lemma mono_bag_of: "mono (bag_of :: 'a list => ('a::order) multiset)"
 | 
|
52  | 
apply (rule monoI)  | 
|
53  | 
apply (unfold prefix_def)  | 
|
54  | 
apply (erule genPrefix.induct, auto)  | 
|
55  | 
apply (erule order_trans)  | 
|
| 35274 | 56  | 
apply simp  | 
| 13798 | 57  | 
done  | 
58  | 
||
59  | 
(** setsum **)  | 
|
60  | 
||
61  | 
declare setsum_cong [cong]  | 
|
62  | 
||
63  | 
lemma bag_of_sublist_lemma:  | 
|
| 15074 | 64  | 
     "(\<Sum>i\<in> A Int lessThan k. {#if i<k then f i else g i#}) =  
 | 
65  | 
      (\<Sum>i\<in> A Int lessThan k. {#f i#})"
 | 
|
| 14114 | 66  | 
by (rule setsum_cong, auto)  | 
| 13798 | 67  | 
|
68  | 
lemma bag_of_sublist:  | 
|
69  | 
"bag_of (sublist l A) =  | 
|
| 15074 | 70  | 
      (\<Sum>i\<in> A Int lessThan (length l). {# l!i #})"
 | 
| 13798 | 71  | 
apply (rule_tac xs = l in rev_induct, simp)  | 
72  | 
apply (simp add: sublist_append Int_insert_right lessThan_Suc nth_append  | 
|
| 14738 | 73  | 
bag_of_sublist_lemma add_ac)  | 
| 13798 | 74  | 
done  | 
75  | 
||
76  | 
||
77  | 
lemma bag_of_sublist_Un_Int:  | 
|
78  | 
"bag_of (sublist l (A Un B)) + bag_of (sublist l (A Int B)) =  | 
|
79  | 
bag_of (sublist l A) + bag_of (sublist l B)"  | 
|
| 15045 | 80  | 
apply (subgoal_tac "A Int B Int {..<length l} =
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
24147 
diff
changeset
 | 
81  | 
                    (A Int {..<length l}) Int (B Int {..<length l}) ")
 | 
| 13798 | 82  | 
apply (simp add: bag_of_sublist Int_Un_distrib2 setsum_Un_Int, blast)  | 
83  | 
done  | 
|
84  | 
||
85  | 
lemma bag_of_sublist_Un_disjoint:  | 
|
86  | 
     "A Int B = {}  
 | 
|
87  | 
==> bag_of (sublist l (A Un B)) =  | 
|
88  | 
bag_of (sublist l A) + bag_of (sublist l B)"  | 
|
| 14114 | 89  | 
by (simp add: bag_of_sublist_Un_Int [symmetric])  | 
| 13798 | 90  | 
|
91  | 
lemma bag_of_sublist_UN_disjoint [rule_format]:  | 
|
92  | 
     "[| finite I; ALL i:I. ALL j:I. i~=j --> A i Int A j = {} |]  
 | 
|
93  | 
==> bag_of (sublist l (UNION I A)) =  | 
|
| 15074 | 94  | 
(\<Sum>i\<in>I. bag_of (sublist l (A i)))"  | 
| 13798 | 95  | 
apply (simp del: UN_simps  | 
96  | 
add: UN_simps [symmetric] add: bag_of_sublist)  | 
|
97  | 
apply (subst setsum_UN_disjoint, auto)  | 
|
98  | 
done  | 
|
99  | 
||
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
100  | 
end  |