author | wenzelm |
Tue, 20 Feb 2018 16:20:36 +0100 | |
changeset 67679 | 8fd84fe1d60b |
parent 67443 | 3abf6a722518 |
child 76213 | e44d86131648 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
1 |
(* Title: ZF/UNITY/Merge.thy |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
2 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
3 |
Copyright 2002 University of Cambridge |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
4 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
5 |
A multiple-client allocator from a single-client allocator: Merge |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
6 |
specification. |
14053
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 Merge imports AllocBase Follows Guar GenPrefix begin |
14073 | 10 |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
11 |
(** Merge specification (the number of inputs is Nclients) ***) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
12 |
(** Parameter A represents the type of items to Merge **) |
14073 | 13 |
|
24893 | 14 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
15 |
(*spec (10)*) |
24893 | 16 |
merge_increasing :: "[i, i, i] =>i" where |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
17 |
"merge_increasing(A, Out, iOut) == program guarantees |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
18 |
(lift(Out) IncreasingWrt prefix(A)/list(A)) Int |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
19 |
(lift(iOut) IncreasingWrt prefix(nat)/list(nat))" |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
20 |
|
24893 | 21 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
22 |
(*spec (11)*) |
24893 | 23 |
merge_eq_Out :: "[i, i] =>i" where |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
24 |
"merge_eq_Out(Out, iOut) == program guarantees |
14073 | 25 |
Always({s \<in> state. length(s`Out) = length(s`iOut)})" |
14053
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 |
(*spec (12)*) |
24893 | 29 |
merge_bounded :: "i=>i" where |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
30 |
"merge_bounded(iOut) == program guarantees |
14073 | 31 |
Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
32 |
|
24893 | 33 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
34 |
(*spec (13)*) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
35 |
(* Parameter A represents the type of tokens *) |
24893 | 36 |
merge_follows :: "[i, i=>i, i, i] =>i" where |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
37 |
"merge_follows(A, In, Out, iOut) == |
14073 | 38 |
(\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A)) |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
39 |
guarantees |
14073 | 40 |
(\<Inter>n \<in> Nclients. |
41 |
(%s. sublist(s`Out, {k \<in> nat. k < length(s`iOut) & |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
42 |
nth(k, s`iOut) = n})) Fols lift(In(n)) |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
43 |
Wrt prefix(A)/list(A))" |
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
44 |
|
24893 | 45 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
46 |
(*spec: preserves part*) |
24893 | 47 |
merge_preserves :: "[i=>i] =>i" where |
14073 | 48 |
"merge_preserves(In) == \<Inter>n \<in> nat. preserves(lift(In(n)))" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
49 |
|
24893 | 50 |
definition |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
51 |
(* environmental constraints*) |
24893 | 52 |
merge_allowed_acts :: "[i, i] =>i" where |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
53 |
"merge_allowed_acts(Out, iOut) == |
14073 | 54 |
{F \<in> program. AllowedActs(F) = |
55 |
cons(id(state), (\<Union>G \<in> preserves(lift(Out)) \<inter> |
|
14057 | 56 |
preserves(lift(iOut)). Acts(G)))}" |
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
57 |
|
24893 | 58 |
definition |
59 |
merge_spec :: "[i, i =>i, i, i]=>i" where |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
60 |
"merge_spec(A, In, Out, iOut) == |
14073 | 61 |
merge_increasing(A, Out, iOut) \<inter> merge_eq_Out(Out, iOut) \<inter> |
62 |
merge_bounded(iOut) \<inter> merge_follows(A, In, Out, iOut) |
|
63 |
\<inter> merge_allowed_acts(Out, iOut) \<inter> merge_preserves(In)" |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
64 |
|
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
65 |
(** State definitions. OUTPUT variables are locals **) |
14073 | 66 |
locale merge = |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
67 |
fixes In \<comment> \<open>merge's INPUT histories: streams to merge\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
68 |
and Out \<comment> \<open>merge's OUTPUT history: merged items\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
69 |
and iOut \<comment> \<open>merge's OUTPUT history: origins of merged items\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
70 |
and A \<comment> \<open>the type of items being merged\<close> |
14073 | 71 |
and M |
72 |
assumes var_assumes [simp]: |
|
73 |
"(\<forall>n. In(n):var) & Out \<in> var & iOut \<in> var" |
|
74 |
and all_distinct_vars: |
|
75 |
"\<forall>n. all_distinct([In(n), Out, iOut])" |
|
76 |
and type_assumes [simp]: |
|
77 |
"(\<forall>n. type_of(In(n))=list(A)) & |
|
78 |
type_of(Out)=list(A) & |
|
79 |
type_of(iOut)=list(nat)" |
|
80 |
and default_val_assumes [simp]: |
|
81 |
"(\<forall>n. default_val(In(n))=Nil) & |
|
82 |
default_val(Out)=Nil & |
|
83 |
default_val(iOut)=Nil" |
|
84 |
and merge_spec: "M \<in> merge_spec(A, In, Out, iOut)" |
|
85 |
||
86 |
||
87 |
lemma (in merge) In_value_type [TC,simp]: "s \<in> state ==> s`In(n) \<in> list(A)" |
|
88 |
apply (unfold state_def) |
|
89 |
apply (drule_tac a = "In (n)" in apply_type) |
|
90 |
apply auto |
|
91 |
done |
|
92 |
||
93 |
lemma (in merge) Out_value_type [TC,simp]: "s \<in> state ==> s`Out \<in> list(A)" |
|
94 |
apply (unfold state_def) |
|
14076 | 95 |
apply (drule_tac a = Out in apply_type, auto) |
14073 | 96 |
done |
97 |
||
98 |
lemma (in merge) iOut_value_type [TC,simp]: "s \<in> state ==> s`iOut \<in> list(nat)" |
|
99 |
apply (unfold state_def) |
|
14076 | 100 |
apply (drule_tac a = iOut in apply_type, auto) |
14073 | 101 |
done |
102 |
||
103 |
lemma (in merge) M_in_program [intro,simp]: "M \<in> program" |
|
104 |
apply (cut_tac merge_spec) |
|
105 |
apply (auto dest: guarantees_type [THEN subsetD] |
|
106 |
simp add: merge_spec_def merge_increasing_def) |
|
107 |
done |
|
108 |
||
109 |
lemma (in merge) merge_Allowed: |
|
46823 | 110 |
"Allowed(M) = (preserves(lift(Out)) \<inter> preserves(lift(iOut)))" |
14073 | 111 |
apply (insert merge_spec preserves_type [of "lift (Out)"]) |
112 |
apply (auto simp add: merge_spec_def merge_allowed_acts_def Allowed_def safety_prop_Acts_iff) |
|
113 |
done |
|
114 |
||
115 |
lemma (in merge) M_ok_iff: |
|
116 |
"G \<in> program ==> |
|
46823 | 117 |
M ok G \<longleftrightarrow> (G \<in> preserves(lift(Out)) & |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
118 |
G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))" |
14073 | 119 |
apply (cut_tac merge_spec) |
120 |
apply (auto simp add: merge_Allowed ok_iff_Allowed) |
|
121 |
done |
|
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
122 |
|
14073 | 123 |
lemma (in merge) merge_Always_Out_eq_iOut: |
124 |
"[| G \<in> preserves(lift(Out)); G \<in> preserves(lift(iOut)); |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
125 |
M \<in> Allowed(G) |] |
14073 | 126 |
==> M \<squnion> G \<in> Always({s \<in> state. length(s`Out)=length(s`iOut)})" |
127 |
apply (frule preserves_type [THEN subsetD]) |
|
128 |
apply (subgoal_tac "G \<in> program") |
|
14076 | 129 |
prefer 2 apply assumption |
14073 | 130 |
apply (frule M_ok_iff) |
131 |
apply (cut_tac merge_spec) |
|
132 |
apply (force dest: guaranteesD simp add: merge_spec_def merge_eq_Out_def) |
|
133 |
done |
|
134 |
||
135 |
lemma (in merge) merge_Bounded: |
|
136 |
"[| G \<in> preserves(lift(iOut)); G \<in> preserves(lift(Out)); |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
137 |
M \<in> Allowed(G) |] ==> |
14073 | 138 |
M \<squnion> G: Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})" |
139 |
apply (frule preserves_type [THEN subsetD]) |
|
140 |
apply (frule M_ok_iff) |
|
141 |
apply (cut_tac merge_spec) |
|
142 |
apply (force dest: guaranteesD simp add: merge_spec_def merge_bounded_def) |
|
143 |
done |
|
144 |
||
145 |
lemma (in merge) merge_bag_Follows_lemma: |
|
146 |
"[| G \<in> preserves(lift(iOut)); |
|
147 |
G: preserves(lift(Out)); M \<in> Allowed(G) |] |
|
148 |
==> M \<squnion> G \<in> Always |
|
149 |
({s \<in> state. msetsum(%i. bag_of(sublist(s`Out, |
|
150 |
{k \<in> nat. k < length(s`iOut) & nth(k, s`iOut)=i})), |
|
151 |
Nclients, A) = bag_of(s`Out)})" |
|
152 |
apply (rule Always_Diff_Un_eq [THEN iffD1]) |
|
153 |
apply (rule_tac [2] state_AlwaysI [THEN Always_weaken]) |
|
14076 | 154 |
apply (rule Always_Int_I [OF merge_Always_Out_eq_iOut merge_Bounded], auto) |
14073 | 155 |
apply (subst bag_of_sublist_UN_disjoint [symmetric]) |
156 |
apply (auto simp add: nat_into_Finite set_of_list_conv_nth [OF iOut_value_type]) |
|
157 |
apply (subgoal_tac " (\<Union>i \<in> Nclients. {k \<in> nat. k < length (x`iOut) & nth (k, x`iOut) = i}) = length (x`iOut) ") |
|
158 |
apply (auto simp add: sublist_upt_eq_take [OF Out_value_type] |
|
159 |
length_type [OF iOut_value_type] |
|
160 |
take_all [OF _ Out_value_type] |
|
161 |
length_type [OF iOut_value_type]) |
|
162 |
apply (rule equalityI) |
|
14076 | 163 |
apply (blast dest: ltD, clarify) |
14073 | 164 |
apply (subgoal_tac "length (x ` iOut) \<in> nat") |
14076 | 165 |
prefer 2 apply (simp add: length_type [OF iOut_value_type]) |
14073 | 166 |
apply (subgoal_tac "xa \<in> nat") |
167 |
apply (simp_all add: Ord_mem_iff_lt) |
|
168 |
prefer 2 apply (blast intro: lt_trans) |
|
59788 | 169 |
apply (drule_tac x = "nth (xa, x`iOut)" and P = "%elt. X (elt) \<longrightarrow> elt<Nclients" for X in bspec) |
14073 | 170 |
apply (simp add: ltI nat_into_Ord) |
171 |
apply (blast dest: ltD) |
|
172 |
done |
|
173 |
||
174 |
theorem (in merge) merge_bag_Follows: |
|
175 |
"M \<in> (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A)) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
176 |
guarantees |
14073 | 177 |
(%s. bag_of(s`Out)) Fols |
178 |
(%s. msetsum(%i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)" |
|
179 |
apply (cut_tac merge_spec) |
|
180 |
apply (rule merge_bag_Follows_lemma [THEN Always_Follows1, THEN guaranteesI]) |
|
14076 | 181 |
apply (simp_all add: M_ok_iff, clarify) |
14073 | 182 |
apply (rule Follows_state_ofD1 [OF Follows_msetsum_UN]) |
183 |
apply (simp_all add: nat_into_Finite bag_of_multiset [of _ A]) |
|
14076 | 184 |
apply (simp add: INT_iff merge_spec_def merge_follows_def, clarify) |
14073 | 185 |
apply (cut_tac merge_spec) |
186 |
apply (subgoal_tac "M ok G") |
|
187 |
prefer 2 apply (force intro: M_ok_iff [THEN iffD2]) |
|
14076 | 188 |
apply (drule guaranteesD, assumption) |
189 |
apply (simp add: merge_spec_def merge_follows_def, blast) |
|
14073 | 190 |
apply (simp cong add: Follows_cong |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
191 |
add: refl_prefix |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
192 |
mono_bag_of [THEN subset_Follows_comp, THEN subsetD, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
193 |
unfolded metacomp_def]) |
14073 | 194 |
done |
195 |
||
14053
4daa384f4fd7
Introduction of the theories UNITY/Merge, UNITY/ClientImpl
paulson
parents:
diff
changeset
|
196 |
end |