| author | hoelzl |
| Thu, 23 Jul 2015 16:40:47 +0200 | |
| changeset 60772 | a0cfa9050fa8 |
| parent 60515 | 484559628038 |
| child 62430 | 9527ff088c15 |
| 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 |
|
| 58889 | 6 |
section{*Common Declarations for Chandy and Charpentier's Allocator*}
|
| 13798 | 7 |
|
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
8 |
theory AllocBase imports "../UNITY_Main" "~~/src/HOL/Library/Multiset_Order" begin |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
9 |
|
| 45827 | 10 |
consts Nclients :: nat (*Number of clients*) |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
11 |
|
| 45827 | 12 |
axiomatization NbT :: nat (*Number of tokens in system*) |
13 |
where NbT_pos: "0 < NbT" |
|
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
14 |
|
|
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
15 |
(*This function merely sums the elements of a list*) |
| 39246 | 16 |
primrec tokens :: "nat list => nat" where |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
17 |
"tokens [] = 0" |
| 39246 | 18 |
| "tokens (x#xs) = x + tokens xs" |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
19 |
|
| 60515 | 20 |
abbreviation (input) "bag_of \<equiv> mset" |
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
21 |
|
| 13798 | 22 |
lemma setsum_fun_mono [rule_format]: |
23 |
"!!f :: nat=>nat. |
|
24 |
(ALL i. i<n --> f i <= g i) --> |
|
25 |
setsum f (lessThan n) <= setsum g (lessThan n)" |
|
26 |
apply (induct_tac "n") |
|
27 |
apply (auto simp add: lessThan_Suc) |
|
28 |
done |
|
29 |
||
| 46911 | 30 |
lemma tokens_mono_prefix: "xs <= ys ==> tokens xs <= tokens ys" |
31 |
by (induct ys arbitrary: xs) (auto simp add: prefix_Cons) |
|
| 13798 | 32 |
|
33 |
lemma mono_tokens: "mono tokens" |
|
34 |
apply (unfold mono_def) |
|
35 |
apply (blast intro: tokens_mono_prefix) |
|
36 |
done |
|
37 |
||
38 |
||
39 |
(** bag_of **) |
|
40 |
||
41 |
lemma bag_of_append [simp]: "bag_of (l@l') = bag_of l + bag_of l'" |
|
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57418
diff
changeset
|
42 |
by (induct l) (simp_all add: ac_simps) |
| 13798 | 43 |
|
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
44 |
lemma subseteq_le_multiset: "A #<=# A + B" |
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
45 |
unfolding le_multiset_def apply (cases B; simp) |
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
46 |
apply (rule HOL.disjI1) |
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
47 |
apply (rule union_less_mono2[of "{#}" "_ + {#_#}" A, simplified])
|
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
48 |
apply (simp add: less_multiset\<^sub>H\<^sub>O) |
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
49 |
done |
|
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
50 |
|
| 13798 | 51 |
lemma mono_bag_of: "mono (bag_of :: 'a list => ('a::order) multiset)"
|
52 |
apply (rule monoI) |
|
53 |
apply (unfold prefix_def) |
|
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
54 |
apply (erule genPrefix.induct, simp_all add: add_right_mono) |
| 13798 | 55 |
apply (erule order_trans) |
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
58889
diff
changeset
|
56 |
apply (simp add: less_eq_multiset_def subseteq_le_multiset) |
| 13798 | 57 |
done |
58 |
||
59 |
(** setsum **) |
|
60 |
||
| 57418 | 61 |
declare setsum.cong [cong] |
| 13798 | 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#})"
|
|
| 57418 | 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 |
|
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57418
diff
changeset
|
73 |
bag_of_sublist_lemma ac_simps) |
| 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}) ")
|
| 57418 | 82 |
apply (simp add: bag_of_sublist Int_Un_distrib2 setsum.union_inter, blast) |
| 13798 | 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) |
|
| 57418 | 97 |
apply (subst setsum.UNION_disjoint, auto) |
| 13798 | 98 |
done |
99 |
||
|
11194
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
100 |
end |