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