11479
|
1 |
(* Title: ZF/UNITY/UNITY.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Sidi O Ehmety, Computer Laboratory
|
|
4 |
Copyright 2001 University of Cambridge
|
|
5 |
|
|
6 |
The basic UNITY theory (revised version, based upon the "co" operator)
|
|
7 |
From Misra, "A Logic for Concurrent Programming", 1994
|
|
8 |
|
|
9 |
Theory ported from HOL.
|
|
10 |
*)
|
|
11 |
|
14077
|
12 |
header {*The Basic UNITY Theory*}
|
|
13 |
|
|
14 |
theory UNITY = State:
|
11479
|
15 |
consts
|
|
16 |
constrains :: "[i, i] => i" (infixl "co" 60)
|
|
17 |
op_unless :: "[i, i] => i" (infixl "unless" 60)
|
|
18 |
|
|
19 |
constdefs
|
14077
|
20 |
program :: i
|
12195
|
21 |
"program == {<init, acts, allowed>:
|
14077
|
22 |
Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
|
|
23 |
id(state) \<in> acts & id(state) \<in> allowed}"
|
11479
|
24 |
|
14077
|
25 |
mk_program :: "[i,i,i]=>i"
|
|
26 |
--{* The definition yields a program thanks to the coercions
|
|
27 |
init \<inter> state, acts \<inter> Pow(state*state), etc. *}
|
11479
|
28 |
"mk_program(init, acts, allowed) ==
|
14077
|
29 |
<init \<inter> state, cons(id(state), acts \<inter> Pow(state*state)),
|
|
30 |
cons(id(state), allowed \<inter> Pow(state*state))>"
|
11479
|
31 |
|
|
32 |
SKIP :: i
|
12195
|
33 |
"SKIP == mk_program(state, 0, Pow(state*state))"
|
11479
|
34 |
|
12195
|
35 |
(* Coercion from anything to program *)
|
14077
|
36 |
programify :: "i=>i"
|
|
37 |
"programify(F) == if F \<in> program then F else SKIP"
|
12195
|
38 |
|
14077
|
39 |
RawInit :: "i=>i"
|
11479
|
40 |
"RawInit(F) == fst(F)"
|
14077
|
41 |
|
|
42 |
Init :: "i=>i"
|
11479
|
43 |
"Init(F) == RawInit(programify(F))"
|
|
44 |
|
14077
|
45 |
RawActs :: "i=>i"
|
12195
|
46 |
"RawActs(F) == cons(id(state), fst(snd(F)))"
|
11479
|
47 |
|
14077
|
48 |
Acts :: "i=>i"
|
11479
|
49 |
"Acts(F) == RawActs(programify(F))"
|
|
50 |
|
14077
|
51 |
RawAllowedActs :: "i=>i"
|
12195
|
52 |
"RawAllowedActs(F) == cons(id(state), snd(snd(F)))"
|
11479
|
53 |
|
14077
|
54 |
AllowedActs :: "i=>i"
|
11479
|
55 |
"AllowedActs(F) == RawAllowedActs(programify(F))"
|
|
56 |
|
14077
|
57 |
|
|
58 |
Allowed :: "i =>i"
|
|
59 |
"Allowed(F) == {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}"
|
11479
|
60 |
|
14077
|
61 |
initially :: "i=>i"
|
|
62 |
"initially(A) == {F \<in> program. Init(F)\<subseteq>A}"
|
|
63 |
|
|
64 |
stable :: "i=>i"
|
11479
|
65 |
"stable(A) == A co A"
|
|
66 |
|
14077
|
67 |
strongest_rhs :: "[i, i] => i"
|
|
68 |
"strongest_rhs(F, A) == Inter({B \<in> Pow(state). F \<in> A co B})"
|
11479
|
69 |
|
14077
|
70 |
invariant :: "i => i"
|
|
71 |
"invariant(A) == initially(A) \<inter> stable(A)"
|
|
72 |
|
14046
|
73 |
(* meta-function composition *)
|
14077
|
74 |
metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl "comp" 65)
|
14046
|
75 |
"f comp g == %x. f(g(x))"
|
11479
|
76 |
|
14046
|
77 |
pg_compl :: "i=>i"
|
|
78 |
"pg_compl(X)== program - X"
|
14077
|
79 |
|
11479
|
80 |
defs
|
14077
|
81 |
constrains_def:
|
|
82 |
"A co B == {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) & st_set(A)}"
|
|
83 |
--{* the condition @{term "st_set(A)"} makes the definition slightly
|
|
84 |
stronger than the HOL one *}
|
|
85 |
|
|
86 |
unless_def: "A unless B == (A - B) co (A Un B)"
|
|
87 |
|
|
88 |
|
|
89 |
(** SKIP **)
|
|
90 |
lemma SKIP_in_program [iff,TC]: "SKIP \<in> program"
|
|
91 |
by (force simp add: SKIP_def program_def mk_program_def)
|
|
92 |
|
|
93 |
|
|
94 |
subsection{*The function @{term programify}, the coercion from anything to
|
|
95 |
program*}
|
|
96 |
|
|
97 |
lemma programify_program [simp]: "F \<in> program ==> programify(F)=F"
|
|
98 |
by (force simp add: programify_def)
|
|
99 |
|
|
100 |
lemma programify_in_program [iff,TC]: "programify(F) \<in> program"
|
|
101 |
by (force simp add: programify_def)
|
|
102 |
|
|
103 |
(** Collapsing rules: to remove programify from expressions **)
|
|
104 |
lemma programify_idem [simp]: "programify(programify(F))=programify(F)"
|
|
105 |
by (force simp add: programify_def)
|
|
106 |
|
|
107 |
lemma Init_programify [simp]: "Init(programify(F)) = Init(F)"
|
|
108 |
by (simp add: Init_def)
|
|
109 |
|
|
110 |
lemma Acts_programify [simp]: "Acts(programify(F)) = Acts(F)"
|
|
111 |
by (simp add: Acts_def)
|
|
112 |
|
|
113 |
lemma AllowedActs_programify [simp]:
|
|
114 |
"AllowedActs(programify(F)) = AllowedActs(F)"
|
|
115 |
by (simp add: AllowedActs_def)
|
|
116 |
|
|
117 |
subsection{*The Inspectors for Programs*}
|
|
118 |
|
|
119 |
lemma id_in_RawActs: "F \<in> program ==>id(state) \<in> RawActs(F)"
|
|
120 |
by (auto simp add: program_def RawActs_def)
|
|
121 |
|
|
122 |
lemma id_in_Acts [iff,TC]: "id(state) \<in> Acts(F)"
|
|
123 |
by (simp add: id_in_RawActs Acts_def)
|
|
124 |
|
|
125 |
lemma id_in_RawAllowedActs: "F \<in> program ==>id(state) \<in> RawAllowedActs(F)"
|
|
126 |
by (auto simp add: program_def RawAllowedActs_def)
|
|
127 |
|
|
128 |
lemma id_in_AllowedActs [iff,TC]: "id(state) \<in> AllowedActs(F)"
|
|
129 |
by (simp add: id_in_RawAllowedActs AllowedActs_def)
|
|
130 |
|
|
131 |
lemma cons_id_Acts [simp]: "cons(id(state), Acts(F)) = Acts(F)"
|
|
132 |
by (simp add: cons_absorb)
|
|
133 |
|
|
134 |
lemma cons_id_AllowedActs [simp]:
|
|
135 |
"cons(id(state), AllowedActs(F)) = AllowedActs(F)"
|
|
136 |
by (simp add: cons_absorb)
|
|
137 |
|
|
138 |
|
|
139 |
subsection{*Types of the Inspectors*}
|
|
140 |
|
|
141 |
lemma RawInit_type: "F \<in> program ==> RawInit(F)\<subseteq>state"
|
|
142 |
by (auto simp add: program_def RawInit_def)
|
|
143 |
|
|
144 |
lemma RawActs_type: "F \<in> program ==> RawActs(F)\<subseteq>Pow(state*state)"
|
|
145 |
by (auto simp add: program_def RawActs_def)
|
|
146 |
|
|
147 |
lemma RawAllowedActs_type:
|
|
148 |
"F \<in> program ==> RawAllowedActs(F)\<subseteq>Pow(state*state)"
|
|
149 |
by (auto simp add: program_def RawAllowedActs_def)
|
|
150 |
|
|
151 |
lemma Init_type: "Init(F)\<subseteq>state"
|
|
152 |
by (simp add: RawInit_type Init_def)
|
|
153 |
|
|
154 |
lemmas InitD = Init_type [THEN subsetD, standard]
|
|
155 |
|
|
156 |
lemma st_set_Init [iff]: "st_set(Init(F))"
|
|
157 |
apply (unfold st_set_def)
|
|
158 |
apply (rule Init_type)
|
|
159 |
done
|
|
160 |
|
|
161 |
lemma Acts_type: "Acts(F)\<subseteq>Pow(state*state)"
|
|
162 |
by (simp add: RawActs_type Acts_def)
|
|
163 |
|
|
164 |
lemma AllowedActs_type: "AllowedActs(F) \<subseteq> Pow(state*state)"
|
|
165 |
by (simp add: RawAllowedActs_type AllowedActs_def)
|
|
166 |
|
|
167 |
(* Needed in Behaviors *)
|
|
168 |
lemma ActsD: "[| act \<in> Acts(F); <s,s'> \<in> act |] ==> s \<in> state & s' \<in> state"
|
|
169 |
by (blast dest: Acts_type [THEN subsetD])
|
|
170 |
|
|
171 |
lemma AllowedActsD:
|
|
172 |
"[| act \<in> AllowedActs(F); <s,s'> \<in> act |] ==> s \<in> state & s' \<in> state"
|
|
173 |
by (blast dest: AllowedActs_type [THEN subsetD])
|
|
174 |
|
|
175 |
subsection{*Simplification rules involving @{term state}, @{term Init},
|
|
176 |
@{term Acts}, and @{term AllowedActs}*}
|
|
177 |
|
|
178 |
text{*But are they really needed?*}
|
|
179 |
|
|
180 |
lemma state_subset_is_Init_iff [iff]: "state \<subseteq> Init(F) <-> Init(F)=state"
|
|
181 |
by (cut_tac F = F in Init_type, auto)
|
|
182 |
|
|
183 |
lemma Pow_state_times_state_is_subset_Acts_iff [iff]:
|
|
184 |
"Pow(state*state) \<subseteq> Acts(F) <-> Acts(F)=Pow(state*state)"
|
|
185 |
by (cut_tac F = F in Acts_type, auto)
|
|
186 |
|
|
187 |
lemma Pow_state_times_state_is_subset_AllowedActs_iff [iff]:
|
|
188 |
"Pow(state*state) \<subseteq> AllowedActs(F) <-> AllowedActs(F)=Pow(state*state)"
|
|
189 |
by (cut_tac F = F in AllowedActs_type, auto)
|
|
190 |
|
|
191 |
subsubsection{*Eliminating @{text "\<inter> state"} from expressions*}
|
|
192 |
|
|
193 |
lemma Init_Int_state [simp]: "Init(F) \<inter> state = Init(F)"
|
|
194 |
by (cut_tac F = F in Init_type, blast)
|
|
195 |
|
|
196 |
lemma state_Int_Init [simp]: "state \<inter> Init(F) = Init(F)"
|
|
197 |
by (cut_tac F = F in Init_type, blast)
|
|
198 |
|
|
199 |
lemma Acts_Int_Pow_state_times_state [simp]:
|
|
200 |
"Acts(F) \<inter> Pow(state*state) = Acts(F)"
|
|
201 |
by (cut_tac F = F in Acts_type, blast)
|
|
202 |
|
|
203 |
lemma state_times_state_Int_Acts [simp]:
|
|
204 |
"Pow(state*state) \<inter> Acts(F) = Acts(F)"
|
|
205 |
by (cut_tac F = F in Acts_type, blast)
|
|
206 |
|
|
207 |
lemma AllowedActs_Int_Pow_state_times_state [simp]:
|
|
208 |
"AllowedActs(F) \<inter> Pow(state*state) = AllowedActs(F)"
|
|
209 |
by (cut_tac F = F in AllowedActs_type, blast)
|
|
210 |
|
|
211 |
lemma state_times_state_Int_AllowedActs [simp]:
|
|
212 |
"Pow(state*state) \<inter> AllowedActs(F) = AllowedActs(F)"
|
|
213 |
by (cut_tac F = F in AllowedActs_type, blast)
|
|
214 |
|
|
215 |
|
|
216 |
subsubsection{*The Opoerator @{term mk_program}*}
|
|
217 |
|
|
218 |
lemma mk_program_in_program [iff,TC]:
|
|
219 |
"mk_program(init, acts, allowed) \<in> program"
|
|
220 |
by (auto simp add: mk_program_def program_def)
|
|
221 |
|
|
222 |
lemma RawInit_eq [simp]:
|
|
223 |
"RawInit(mk_program(init, acts, allowed)) = init \<inter> state"
|
|
224 |
by (auto simp add: mk_program_def RawInit_def)
|
|
225 |
|
|
226 |
lemma RawActs_eq [simp]:
|
|
227 |
"RawActs(mk_program(init, acts, allowed)) =
|
|
228 |
cons(id(state), acts \<inter> Pow(state*state))"
|
|
229 |
by (auto simp add: mk_program_def RawActs_def)
|
|
230 |
|
|
231 |
lemma RawAllowedActs_eq [simp]:
|
|
232 |
"RawAllowedActs(mk_program(init, acts, allowed)) =
|
|
233 |
cons(id(state), allowed \<inter> Pow(state*state))"
|
|
234 |
by (auto simp add: mk_program_def RawAllowedActs_def)
|
|
235 |
|
|
236 |
lemma Init_eq [simp]: "Init(mk_program(init, acts, allowed)) = init \<inter> state"
|
|
237 |
by (simp add: Init_def)
|
|
238 |
|
|
239 |
lemma Acts_eq [simp]:
|
|
240 |
"Acts(mk_program(init, acts, allowed)) =
|
|
241 |
cons(id(state), acts \<inter> Pow(state*state))"
|
|
242 |
by (simp add: Acts_def)
|
|
243 |
|
|
244 |
lemma AllowedActs_eq [simp]:
|
|
245 |
"AllowedActs(mk_program(init, acts, allowed))=
|
|
246 |
cons(id(state), allowed \<inter> Pow(state*state))"
|
|
247 |
by (simp add: AllowedActs_def)
|
|
248 |
|
|
249 |
(**Init, Acts, and AlowedActs of SKIP **)
|
|
250 |
|
|
251 |
lemma RawInit_SKIP [simp]: "RawInit(SKIP) = state"
|
|
252 |
by (simp add: SKIP_def)
|
|
253 |
|
|
254 |
lemma RawAllowedActs_SKIP [simp]: "RawAllowedActs(SKIP) = Pow(state*state)"
|
|
255 |
by (force simp add: SKIP_def)
|
|
256 |
|
|
257 |
lemma RawActs_SKIP [simp]: "RawActs(SKIP) = {id(state)}"
|
|
258 |
by (force simp add: SKIP_def)
|
|
259 |
|
|
260 |
lemma Init_SKIP [simp]: "Init(SKIP) = state"
|
|
261 |
by (force simp add: SKIP_def)
|
|
262 |
|
|
263 |
lemma Acts_SKIP [simp]: "Acts(SKIP) = {id(state)}"
|
|
264 |
by (force simp add: SKIP_def)
|
|
265 |
|
|
266 |
lemma AllowedActs_SKIP [simp]: "AllowedActs(SKIP) = Pow(state*state)"
|
|
267 |
by (force simp add: SKIP_def)
|
|
268 |
|
|
269 |
(** Equality of UNITY programs **)
|
|
270 |
|
|
271 |
lemma raw_surjective_mk_program:
|
|
272 |
"F \<in> program ==> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F"
|
|
273 |
apply (auto simp add: program_def mk_program_def RawInit_def RawActs_def
|
|
274 |
RawAllowedActs_def, blast+)
|
|
275 |
done
|
|
276 |
|
|
277 |
lemma surjective_mk_program [simp]:
|
|
278 |
"mk_program(Init(F), Acts(F), AllowedActs(F)) = programify(F)"
|
|
279 |
by (auto simp add: raw_surjective_mk_program Init_def Acts_def AllowedActs_def)
|
|
280 |
|
|
281 |
lemma program_equalityI:
|
|
282 |
"[|Init(F) = Init(G); Acts(F) = Acts(G);
|
|
283 |
AllowedActs(F) = AllowedActs(G); F \<in> program; G \<in> program |] ==> F = G"
|
|
284 |
apply (subgoal_tac "programify(F) = programify(G)")
|
|
285 |
apply simp
|
|
286 |
apply (simp only: surjective_mk_program [symmetric])
|
|
287 |
done
|
|
288 |
|
|
289 |
lemma program_equalityE:
|
|
290 |
"[|F = G;
|
|
291 |
[|Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G) |]
|
|
292 |
==> P |]
|
|
293 |
==> P"
|
|
294 |
by force
|
|
295 |
|
|
296 |
|
|
297 |
lemma program_equality_iff:
|
|
298 |
"[| F \<in> program; G \<in> program |] ==>(F=G) <->
|
|
299 |
(Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(G))"
|
|
300 |
by (blast intro: program_equalityI program_equalityE)
|
|
301 |
|
|
302 |
subsection{*These rules allow "lazy" definition expansion*}
|
|
303 |
|
|
304 |
lemma def_prg_Init:
|
|
305 |
"F == mk_program (init,acts,allowed) ==> Init(F) = init \<inter> state"
|
|
306 |
by auto
|
|
307 |
|
|
308 |
lemma def_prg_Acts:
|
|
309 |
"F == mk_program (init,acts,allowed)
|
|
310 |
==> Acts(F) = cons(id(state), acts \<inter> Pow(state*state))"
|
|
311 |
by auto
|
|
312 |
|
|
313 |
lemma def_prg_AllowedActs:
|
|
314 |
"F == mk_program (init,acts,allowed)
|
|
315 |
==> AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
|
|
316 |
by auto
|
|
317 |
|
|
318 |
lemma def_prg_simps:
|
|
319 |
"[| F == mk_program (init,acts,allowed) |]
|
|
320 |
==> Init(F) = init \<inter> state &
|
|
321 |
Acts(F) = cons(id(state), acts \<inter> Pow(state*state)) &
|
|
322 |
AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
|
|
323 |
by auto
|
|
324 |
|
|
325 |
|
|
326 |
(*An action is expanded only if a pair of states is being tested against it*)
|
|
327 |
lemma def_act_simp:
|
|
328 |
"[| act == {<s,s'> \<in> A*B. P(s, s')} |]
|
|
329 |
==> (<s,s'> \<in> act) <-> (<s,s'> \<in> A*B & P(s, s'))"
|
|
330 |
by auto
|
|
331 |
|
|
332 |
(*A set is expanded only if an element is being tested against it*)
|
|
333 |
lemma def_set_simp: "A == B ==> (x \<in> A) <-> (x \<in> B)"
|
|
334 |
by auto
|
|
335 |
|
|
336 |
|
|
337 |
subsection{*The Constrains Operator*}
|
|
338 |
|
|
339 |
lemma constrains_type: "A co B \<subseteq> program"
|
|
340 |
by (force simp add: constrains_def)
|
|
341 |
|
|
342 |
lemma constrainsI:
|
|
343 |
"[|(!!act s s'. [| act: Acts(F); <s,s'> \<in> act; s \<in> A|] ==> s' \<in> A');
|
|
344 |
F \<in> program; st_set(A) |] ==> F \<in> A co A'"
|
|
345 |
by (force simp add: constrains_def)
|
|
346 |
|
|
347 |
lemma constrainsD:
|
|
348 |
"F \<in> A co B ==> \<forall>act \<in> Acts(F). act``A\<subseteq>B"
|
|
349 |
by (force simp add: constrains_def)
|
|
350 |
|
|
351 |
lemma constrainsD2: "F \<in> A co B ==> F \<in> program & st_set(A)"
|
|
352 |
by (force simp add: constrains_def)
|
|
353 |
|
|
354 |
lemma constrains_empty [iff]: "F \<in> 0 co B <-> F \<in> program"
|
|
355 |
by (force simp add: constrains_def st_set_def)
|
|
356 |
|
|
357 |
lemma constrains_empty2 [iff]: "(F \<in> A co 0) <-> (A=0 & F \<in> program)"
|
|
358 |
by (force simp add: constrains_def st_set_def)
|
|
359 |
|
|
360 |
lemma constrains_state [iff]: "(F \<in> state co B) <-> (state\<subseteq>B & F \<in> program)"
|
|
361 |
apply (cut_tac F = F in Acts_type)
|
|
362 |
apply (force simp add: constrains_def st_set_def)
|
|
363 |
done
|
|
364 |
|
|
365 |
lemma constrains_state2 [iff]: "F \<in> A co state <-> (F \<in> program & st_set(A))"
|
|
366 |
apply (cut_tac F = F in Acts_type)
|
|
367 |
apply (force simp add: constrains_def st_set_def)
|
|
368 |
done
|
|
369 |
|
|
370 |
(*monotonic in 2nd argument*)
|
|
371 |
lemma constrains_weaken_R:
|
|
372 |
"[| F \<in> A co A'; A'\<subseteq>B' |] ==> F \<in> A co B'"
|
|
373 |
apply (unfold constrains_def, blast)
|
|
374 |
done
|
|
375 |
|
|
376 |
(*anti-monotonic in 1st argument*)
|
|
377 |
lemma constrains_weaken_L:
|
|
378 |
"[| F \<in> A co A'; B\<subseteq>A |] ==> F \<in> B co A'"
|
|
379 |
apply (unfold constrains_def st_set_def, blast)
|
|
380 |
done
|
|
381 |
|
|
382 |
lemma constrains_weaken:
|
|
383 |
"[| F \<in> A co A'; B\<subseteq>A; A'\<subseteq>B' |] ==> F \<in> B co B'"
|
|
384 |
apply (drule constrains_weaken_R)
|
|
385 |
apply (drule_tac [2] constrains_weaken_L, blast+)
|
|
386 |
done
|
|
387 |
|
|
388 |
|
|
389 |
subsection{*Constrains and Union*}
|
|
390 |
|
|
391 |
lemma constrains_Un:
|
|
392 |
"[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A Un B) co (A' Un B')"
|
|
393 |
by (auto simp add: constrains_def st_set_def, force)
|
|
394 |
|
|
395 |
lemma constrains_UN:
|
|
396 |
"[|!!i. i \<in> I ==> F \<in> A(i) co A'(i); F \<in> program |]
|
|
397 |
==> F \<in> (\<Union>i \<in> I. A(i)) co (\<Union>i \<in> I. A'(i))"
|
|
398 |
by (force simp add: constrains_def st_set_def)
|
|
399 |
|
|
400 |
lemma constrains_Un_distrib:
|
|
401 |
"(A Un B) co C = (A co C) \<inter> (B co C)"
|
|
402 |
by (force simp add: constrains_def st_set_def)
|
|
403 |
|
|
404 |
lemma constrains_UN_distrib:
|
|
405 |
"i \<in> I ==> (\<Union>i \<in> I. A(i)) co B = (\<Inter>i \<in> I. A(i) co B)"
|
|
406 |
by (force simp add: constrains_def st_set_def)
|
|
407 |
|
|
408 |
|
|
409 |
subsection{*Constrains and Intersection*}
|
|
410 |
|
|
411 |
lemma constrains_Int_distrib: "C co (A \<inter> B) = (C co A) \<inter> (C co B)"
|
|
412 |
by (force simp add: constrains_def st_set_def)
|
|
413 |
|
|
414 |
lemma constrains_INT_distrib:
|
|
415 |
"x \<in> I ==> A co (\<Inter>i \<in> I. B(i)) = (\<Inter>i \<in> I. A co B(i))"
|
|
416 |
by (force simp add: constrains_def st_set_def)
|
|
417 |
|
|
418 |
lemma constrains_Int:
|
|
419 |
"[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<inter> B) co (A' \<inter> B')"
|
|
420 |
by (force simp add: constrains_def st_set_def)
|
|
421 |
|
|
422 |
lemma constrains_INT [rule_format]:
|
|
423 |
"[| \<forall>i \<in> I. F \<in> A(i) co A'(i); F \<in> program|]
|
|
424 |
==> F \<in> (\<Inter>i \<in> I. A(i)) co (\<Inter>i \<in> I. A'(i))"
|
|
425 |
apply (case_tac "I=0")
|
|
426 |
apply (simp add: Inter_def)
|
|
427 |
apply (erule not_emptyE)
|
|
428 |
apply (auto simp add: constrains_def st_set_def, blast)
|
|
429 |
apply (drule bspec, assumption, force)
|
|
430 |
done
|
|
431 |
|
|
432 |
(* The rule below simulates the HOL's one for (\<Inter>z. A i) co (\<Inter>z. B i) *)
|
|
433 |
lemma constrains_All:
|
|
434 |
"[| \<forall>z. F:{s \<in> state. P(s, z)} co {s \<in> state. Q(s, z)}; F \<in> program |]==>
|
|
435 |
F:{s \<in> state. \<forall>z. P(s, z)} co {s \<in> state. \<forall>z. Q(s, z)}"
|
|
436 |
by (unfold constrains_def, blast)
|
|
437 |
|
|
438 |
lemma constrains_imp_subset:
|
|
439 |
"[| F \<in> A co A' |] ==> A \<subseteq> A'"
|
|
440 |
by (unfold constrains_def st_set_def, force)
|
|
441 |
|
|
442 |
(*The reasoning is by subsets since "co" refers to single actions
|
|
443 |
only. So this rule isn't that useful.*)
|
|
444 |
|
|
445 |
lemma constrains_trans: "[| F \<in> A co B; F \<in> B co C |] ==> F \<in> A co C"
|
|
446 |
by (unfold constrains_def st_set_def, auto, blast)
|
|
447 |
|
|
448 |
lemma constrains_cancel:
|
|
449 |
"[| F \<in> A co (A' Un B); F \<in> B co B' |] ==> F \<in> A co (A' Un B')"
|
|
450 |
apply (drule_tac A = B in constrains_imp_subset)
|
|
451 |
apply (blast intro: constrains_weaken_R)
|
|
452 |
done
|
|
453 |
|
|
454 |
|
|
455 |
subsection{*The Unless Operator*}
|
|
456 |
|
|
457 |
lemma unless_type: "A unless B \<subseteq> program"
|
|
458 |
by (force simp add: unless_def constrains_def)
|
|
459 |
|
|
460 |
lemma unlessI: "[| F \<in> (A-B) co (A Un B) |] ==> F \<in> A unless B"
|
|
461 |
apply (unfold unless_def)
|
|
462 |
apply (blast dest: constrainsD2)
|
|
463 |
done
|
|
464 |
|
|
465 |
lemma unlessD: "F :A unless B ==> F \<in> (A-B) co (A Un B)"
|
|
466 |
by (unfold unless_def, auto)
|
|
467 |
|
|
468 |
|
|
469 |
subsection{*The Operator @{term initially}*}
|
|
470 |
|
|
471 |
lemma initially_type: "initially(A) \<subseteq> program"
|
|
472 |
by (unfold initially_def, blast)
|
|
473 |
|
|
474 |
lemma initiallyI: "[| F \<in> program; Init(F)\<subseteq>A |] ==> F \<in> initially(A)"
|
|
475 |
by (unfold initially_def, blast)
|
|
476 |
|
|
477 |
lemma initiallyD: "F \<in> initially(A) ==> Init(F)\<subseteq>A"
|
|
478 |
by (unfold initially_def, blast)
|
|
479 |
|
|
480 |
|
|
481 |
subsection{*The Operator @{term stable}*}
|
|
482 |
|
|
483 |
lemma stable_type: "stable(A)\<subseteq>program"
|
|
484 |
by (unfold stable_def constrains_def, blast)
|
|
485 |
|
|
486 |
lemma stableI: "F \<in> A co A ==> F \<in> stable(A)"
|
|
487 |
by (unfold stable_def, assumption)
|
|
488 |
|
|
489 |
lemma stableD: "F \<in> stable(A) ==> F \<in> A co A"
|
|
490 |
by (unfold stable_def, assumption)
|
|
491 |
|
|
492 |
lemma stableD2: "F \<in> stable(A) ==> F \<in> program & st_set(A)"
|
|
493 |
by (unfold stable_def constrains_def, auto)
|
|
494 |
|
|
495 |
lemma stable_state [simp]: "stable(state) = program"
|
|
496 |
by (auto simp add: stable_def constrains_def dest: Acts_type [THEN subsetD])
|
|
497 |
|
|
498 |
|
|
499 |
lemma stable_unless: "stable(A)= A unless 0"
|
|
500 |
by (auto simp add: unless_def stable_def)
|
|
501 |
|
|
502 |
|
|
503 |
subsection{*Union and Intersection with @{term stable}*}
|
|
504 |
|
|
505 |
lemma stable_Un:
|
|
506 |
"[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable(A Un A')"
|
|
507 |
apply (unfold stable_def)
|
|
508 |
apply (blast intro: constrains_Un)
|
|
509 |
done
|
|
510 |
|
|
511 |
lemma stable_UN:
|
|
512 |
"[|!!i. i\<in>I ==> F \<in> stable(A(i)); F \<in> program |]
|
|
513 |
==> F \<in> stable (\<Union>i \<in> I. A(i))"
|
|
514 |
apply (unfold stable_def)
|
|
515 |
apply (blast intro: constrains_UN)
|
|
516 |
done
|
|
517 |
|
|
518 |
lemma stable_Int:
|
|
519 |
"[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable (A \<inter> A')"
|
|
520 |
apply (unfold stable_def)
|
|
521 |
apply (blast intro: constrains_Int)
|
|
522 |
done
|
|
523 |
|
|
524 |
lemma stable_INT:
|
|
525 |
"[| !!i. i \<in> I ==> F \<in> stable(A(i)); F \<in> program |]
|
|
526 |
==> F \<in> stable (\<Inter>i \<in> I. A(i))"
|
|
527 |
apply (unfold stable_def)
|
|
528 |
apply (blast intro: constrains_INT)
|
|
529 |
done
|
|
530 |
|
|
531 |
lemma stable_All:
|
|
532 |
"[|\<forall>z. F \<in> stable({s \<in> state. P(s, z)}); F \<in> program|]
|
|
533 |
==> F \<in> stable({s \<in> state. \<forall>z. P(s, z)})"
|
|
534 |
apply (unfold stable_def)
|
|
535 |
apply (rule constrains_All, auto)
|
|
536 |
done
|
|
537 |
|
|
538 |
lemma stable_constrains_Un:
|
|
539 |
"[| F \<in> stable(C); F \<in> A co (C Un A') |] ==> F \<in> (C Un A) co (C Un A')"
|
|
540 |
apply (unfold stable_def constrains_def st_set_def, auto)
|
|
541 |
apply (blast dest!: bspec)
|
|
542 |
done
|
|
543 |
|
|
544 |
lemma stable_constrains_Int:
|
|
545 |
"[| F \<in> stable(C); F \<in> (C \<inter> A) co A' |] ==> F \<in> (C \<inter> A) co (C \<inter> A')"
|
|
546 |
by (unfold stable_def constrains_def st_set_def, blast)
|
|
547 |
|
|
548 |
(* [| F \<in> stable(C); F \<in> (C \<inter> A) co A |] ==> F \<in> stable(C \<inter> A) *)
|
|
549 |
lemmas stable_constrains_stable = stable_constrains_Int [THEN stableI, standard]
|
|
550 |
|
|
551 |
subsection{*The Operator @{term invariant}*}
|
|
552 |
|
|
553 |
lemma invariant_type: "invariant(A) \<subseteq> program"
|
|
554 |
apply (unfold invariant_def)
|
|
555 |
apply (blast dest: stable_type [THEN subsetD])
|
|
556 |
done
|
|
557 |
|
|
558 |
lemma invariantI: "[| Init(F)\<subseteq>A; F \<in> stable(A) |] ==> F \<in> invariant(A)"
|
|
559 |
apply (unfold invariant_def initially_def)
|
|
560 |
apply (frule stable_type [THEN subsetD], auto)
|
|
561 |
done
|
|
562 |
|
|
563 |
lemma invariantD: "F \<in> invariant(A) ==> Init(F)\<subseteq>A & F \<in> stable(A)"
|
|
564 |
by (unfold invariant_def initially_def, auto)
|
|
565 |
|
|
566 |
lemma invariantD2: "F \<in> invariant(A) ==> F \<in> program & st_set(A)"
|
|
567 |
apply (unfold invariant_def)
|
|
568 |
apply (blast dest: stableD2)
|
|
569 |
done
|
|
570 |
|
|
571 |
text{*Could also say
|
|
572 |
@{term "invariant(A) \<inter> invariant(B) \<subseteq> invariant (A \<inter> B)"}*}
|
|
573 |
lemma invariant_Int:
|
|
574 |
"[| F \<in> invariant(A); F \<in> invariant(B) |] ==> F \<in> invariant(A \<inter> B)"
|
|
575 |
apply (unfold invariant_def initially_def)
|
|
576 |
apply (simp add: stable_Int, blast)
|
|
577 |
done
|
|
578 |
|
|
579 |
|
|
580 |
subsection{*The Elimination Theorem*}
|
|
581 |
|
|
582 |
(** The "free" m has become universally quantified!
|
|
583 |
Should the premise be !!m instead of \<forall>m ? Would make it harder
|
|
584 |
to use in forward proof. **)
|
|
585 |
|
|
586 |
(* The general case easier to prove that le special case! *)
|
|
587 |
lemma "elimination":
|
|
588 |
"[| \<forall>m \<in> M. F \<in> {s \<in> A. x(s) = m} co B(m); F \<in> program |]
|
|
589 |
==> F \<in> {s \<in> A. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
|
|
590 |
by (auto simp add: constrains_def st_set_def, blast)
|
|
591 |
|
|
592 |
(* As above, but for the special case of A=state *)
|
|
593 |
lemma elimination2:
|
|
594 |
"[| \<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} co B(m); F \<in> program |]
|
|
595 |
==> F:{s \<in> state. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
|
|
596 |
by (rule UNITY.elimination, auto)
|
|
597 |
|
|
598 |
subsection{*The Operator @{term strongest_rhs}*}
|
|
599 |
|
|
600 |
lemma constrains_strongest_rhs:
|
|
601 |
"[| F \<in> program; st_set(A) |] ==> F \<in> A co (strongest_rhs(F,A))"
|
|
602 |
by (auto simp add: constrains_def strongest_rhs_def st_set_def
|
|
603 |
dest: Acts_type [THEN subsetD])
|
|
604 |
|
|
605 |
lemma strongest_rhs_is_strongest:
|
|
606 |
"[| F \<in> A co B; st_set(B) |] ==> strongest_rhs(F,A) \<subseteq> B"
|
|
607 |
by (auto simp add: constrains_def strongest_rhs_def st_set_def)
|
|
608 |
|
|
609 |
ML
|
|
610 |
{*
|
|
611 |
val constrains_def = thm "constrains_def";
|
|
612 |
val stable_def = thm "stable_def";
|
|
613 |
val invariant_def = thm "invariant_def";
|
|
614 |
val unless_def = thm "unless_def";
|
|
615 |
val initially_def = thm "initially_def";
|
|
616 |
val SKIP_def = thm "SKIP_def";
|
|
617 |
val Allowed_def = thm "Allowed_def";
|
|
618 |
val programify_def = thm "programify_def";
|
|
619 |
val metacomp_def = thm "metacomp_def";
|
|
620 |
|
|
621 |
val id_in_Acts = thm "id_in_Acts";
|
|
622 |
val id_in_RawAllowedActs = thm "id_in_RawAllowedActs";
|
|
623 |
val id_in_AllowedActs = thm "id_in_AllowedActs";
|
|
624 |
val cons_id_Acts = thm "cons_id_Acts";
|
|
625 |
val cons_id_AllowedActs = thm "cons_id_AllowedActs";
|
|
626 |
val Init_type = thm "Init_type";
|
|
627 |
val st_set_Init = thm "st_set_Init";
|
|
628 |
val Acts_type = thm "Acts_type";
|
|
629 |
val AllowedActs_type = thm "AllowedActs_type";
|
|
630 |
val ActsD = thm "ActsD";
|
|
631 |
val AllowedActsD = thm "AllowedActsD";
|
|
632 |
val mk_program_in_program = thm "mk_program_in_program";
|
|
633 |
val Init_eq = thm "Init_eq";
|
|
634 |
val Acts_eq = thm "Acts_eq";
|
|
635 |
val AllowedActs_eq = thm "AllowedActs_eq";
|
|
636 |
val Init_SKIP = thm "Init_SKIP";
|
|
637 |
val Acts_SKIP = thm "Acts_SKIP";
|
|
638 |
val AllowedActs_SKIP = thm "AllowedActs_SKIP";
|
|
639 |
val raw_surjective_mk_program = thm "raw_surjective_mk_program";
|
|
640 |
val surjective_mk_program = thm "surjective_mk_program";
|
|
641 |
val program_equalityI = thm "program_equalityI";
|
|
642 |
val program_equalityE = thm "program_equalityE";
|
|
643 |
val program_equality_iff = thm "program_equality_iff";
|
|
644 |
val def_prg_Init = thm "def_prg_Init";
|
|
645 |
val def_prg_Acts = thm "def_prg_Acts";
|
|
646 |
val def_prg_AllowedActs = thm "def_prg_AllowedActs";
|
|
647 |
val def_prg_simps = thm "def_prg_simps";
|
|
648 |
val def_act_simp = thm "def_act_simp";
|
|
649 |
val def_set_simp = thm "def_set_simp";
|
|
650 |
val constrains_type = thm "constrains_type";
|
|
651 |
val constrainsI = thm "constrainsI";
|
|
652 |
val constrainsD = thm "constrainsD";
|
|
653 |
val constrainsD2 = thm "constrainsD2";
|
|
654 |
val constrains_empty = thm "constrains_empty";
|
|
655 |
val constrains_empty2 = thm "constrains_empty2";
|
|
656 |
val constrains_state = thm "constrains_state";
|
|
657 |
val constrains_state2 = thm "constrains_state2";
|
|
658 |
val constrains_weaken_R = thm "constrains_weaken_R";
|
|
659 |
val constrains_weaken_L = thm "constrains_weaken_L";
|
|
660 |
val constrains_weaken = thm "constrains_weaken";
|
|
661 |
val constrains_Un = thm "constrains_Un";
|
|
662 |
val constrains_UN = thm "constrains_UN";
|
|
663 |
val constrains_Un_distrib = thm "constrains_Un_distrib";
|
|
664 |
val constrains_UN_distrib = thm "constrains_UN_distrib";
|
|
665 |
val constrains_Int_distrib = thm "constrains_Int_distrib";
|
|
666 |
val constrains_INT_distrib = thm "constrains_INT_distrib";
|
|
667 |
val constrains_Int = thm "constrains_Int";
|
|
668 |
val constrains_INT = thm "constrains_INT";
|
|
669 |
val constrains_All = thm "constrains_All";
|
|
670 |
val constrains_imp_subset = thm "constrains_imp_subset";
|
|
671 |
val constrains_trans = thm "constrains_trans";
|
|
672 |
val constrains_cancel = thm "constrains_cancel";
|
|
673 |
val unless_type = thm "unless_type";
|
|
674 |
val unlessI = thm "unlessI";
|
|
675 |
val unlessD = thm "unlessD";
|
|
676 |
val initially_type = thm "initially_type";
|
|
677 |
val initiallyI = thm "initiallyI";
|
|
678 |
val initiallyD = thm "initiallyD";
|
|
679 |
val stable_type = thm "stable_type";
|
|
680 |
val stableI = thm "stableI";
|
|
681 |
val stableD = thm "stableD";
|
|
682 |
val stableD2 = thm "stableD2";
|
|
683 |
val stable_state = thm "stable_state";
|
|
684 |
val stable_unless = thm "stable_unless";
|
|
685 |
val stable_Un = thm "stable_Un";
|
|
686 |
val stable_UN = thm "stable_UN";
|
|
687 |
val stable_Int = thm "stable_Int";
|
|
688 |
val stable_INT = thm "stable_INT";
|
|
689 |
val stable_All = thm "stable_All";
|
|
690 |
val stable_constrains_Un = thm "stable_constrains_Un";
|
|
691 |
val stable_constrains_Int = thm "stable_constrains_Int";
|
|
692 |
val invariant_type = thm "invariant_type";
|
|
693 |
val invariantI = thm "invariantI";
|
|
694 |
val invariantD = thm "invariantD";
|
|
695 |
val invariantD2 = thm "invariantD2";
|
|
696 |
val invariant_Int = thm "invariant_Int";
|
|
697 |
val elimination = thm "elimination";
|
|
698 |
val elimination2 = thm "elimination2";
|
|
699 |
val constrains_strongest_rhs = thm "constrains_strongest_rhs";
|
|
700 |
val strongest_rhs_is_strongest = thm "strongest_rhs_is_strongest";
|
|
701 |
|
|
702 |
fun simp_of_act def = def RS def_act_simp;
|
|
703 |
fun simp_of_set def = def RS def_set_simp;
|
|
704 |
*}
|
|
705 |
|
11479
|
706 |
end
|