13020
|
1 |
|
|
2 |
header {* \section{Operational Semantics} *}
|
|
3 |
|
16417
|
4 |
theory OG_Tran imports OG_Com begin
|
13020
|
5 |
|
|
6 |
types
|
|
7 |
'a ann_com_op = "('a ann_com) option"
|
|
8 |
'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
|
|
9 |
|
|
10 |
consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
|
|
11 |
primrec "com (c, q) = c"
|
|
12 |
|
|
13 |
consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
|
|
14 |
primrec "post (c, q) = q"
|
|
15 |
|
|
16 |
constdefs
|
|
17 |
All_None :: "'a ann_triple_op list \<Rightarrow> bool"
|
|
18 |
"All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
|
|
19 |
|
|
20 |
subsection {* The Transition Relation *}
|
|
21 |
|
|
22 |
consts
|
|
23 |
ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"
|
|
24 |
transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
|
|
25 |
|
|
26 |
syntax
|
|
27 |
"_ann_transition" :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
|
|
28 |
("_ -1\<rightarrow> _"[81,81] 100)
|
|
29 |
"_ann_transition_n" :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
|
|
30 |
\<Rightarrow> bool" ("_ -_\<rightarrow> _"[81,81] 100)
|
|
31 |
"_ann_transition_*" :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
|
|
32 |
("_ -*\<rightarrow> _"[81,81] 100)
|
|
33 |
|
|
34 |
"_transition" :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool" ("_ -P1\<rightarrow> _"[81,81] 100)
|
|
35 |
"_transition_n" :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
|
|
36 |
("_ -P_\<rightarrow> _"[81,81,81] 100)
|
|
37 |
"_transition_*" :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool" ("_ -P*\<rightarrow> _"[81,81] 100)
|
|
38 |
|
|
39 |
text {* The corresponding syntax translations are: *}
|
|
40 |
|
|
41 |
translations
|
|
42 |
"con_0 -1\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition"
|
|
43 |
"con_0 -n\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition^n"
|
|
44 |
"con_0 -*\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition\<^sup>*"
|
|
45 |
|
|
46 |
"con_0 -P1\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition"
|
|
47 |
"con_0 -Pn\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition^n"
|
|
48 |
"con_0 -P*\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition\<^sup>*"
|
|
49 |
|
|
50 |
inductive ann_transition transition
|
|
51 |
intros
|
|
52 |
AnnBasic: "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
|
|
53 |
|
|
54 |
AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow>
|
|
55 |
(Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"
|
|
56 |
AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow>
|
|
57 |
(Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"
|
|
58 |
|
|
59 |
AnnCond1T: "s \<in> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"
|
|
60 |
AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"
|
|
61 |
|
|
62 |
AnnCond2T: "s \<in> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"
|
|
63 |
AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"
|
|
64 |
|
|
65 |
AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"
|
|
66 |
AnnWhileT: "s \<in> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow>
|
|
67 |
(Some (AnnSeq c (AnnWhile i b i c)), s)"
|
|
68 |
|
|
69 |
AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
|
|
70 |
(Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)"
|
|
71 |
|
|
72 |
Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
|
|
73 |
\<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
|
|
74 |
|
|
75 |
Basic: "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"
|
|
76 |
|
|
77 |
Seq1: "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"
|
|
78 |
Seq2: "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"
|
|
79 |
|
|
80 |
CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"
|
|
81 |
CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"
|
|
82 |
|
|
83 |
WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"
|
|
84 |
WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"
|
|
85 |
|
|
86 |
monos "rtrancl_mono"
|
|
87 |
|
|
88 |
subsection {* Definition of Semantics *}
|
|
89 |
|
|
90 |
constdefs
|
|
91 |
ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set"
|
|
92 |
"ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
|
|
93 |
|
|
94 |
ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set"
|
|
95 |
"ann_SEM c S \<equiv> \<Union>ann_sem c ` S"
|
|
96 |
|
|
97 |
sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set"
|
|
98 |
"sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
|
|
99 |
|
|
100 |
SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set"
|
|
101 |
"SEM c S \<equiv> \<Union>sem c ` S "
|
|
102 |
|
|
103 |
syntax "_Omega" :: "'a com" ("\<Omega>" 63)
|
|
104 |
translations "\<Omega>" \<rightleftharpoons> "While UNIV UNIV (Basic id)"
|
|
105 |
|
|
106 |
consts fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com"
|
|
107 |
primrec
|
|
108 |
"fwhile b c 0 = \<Omega>"
|
|
109 |
"fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"
|
|
110 |
|
|
111 |
subsubsection {* Proofs *}
|
|
112 |
|
|
113 |
declare ann_transition_transition.intros [intro]
|
|
114 |
inductive_cases transition_cases:
|
|
115 |
"(Parallel T,s) -P1\<rightarrow> t"
|
|
116 |
"(Basic f, s) -P1\<rightarrow> t"
|
|
117 |
"(Seq c1 c2, s) -P1\<rightarrow> t"
|
|
118 |
"(Cond b c1 c2, s) -P1\<rightarrow> t"
|
|
119 |
"(While b i c, s) -P1\<rightarrow> t"
|
|
120 |
|
|
121 |
lemma Parallel_empty_lemma [rule_format (no_asm)]:
|
|
122 |
"(Parallel [],s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=[] \<and> n=0 \<and> s=t"
|
|
123 |
apply(induct n)
|
|
124 |
apply(simp (no_asm))
|
|
125 |
apply clarify
|
|
126 |
apply(drule rel_pow_Suc_D2)
|
|
127 |
apply(force elim:transition_cases)
|
|
128 |
done
|
|
129 |
|
|
130 |
lemma Parallel_AllNone_lemma [rule_format (no_asm)]:
|
|
131 |
"All_None Ss \<longrightarrow> (Parallel Ss,s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=Ss \<and> n=0 \<and> s=t"
|
|
132 |
apply(induct "n")
|
|
133 |
apply(simp (no_asm))
|
|
134 |
apply clarify
|
|
135 |
apply(drule rel_pow_Suc_D2)
|
|
136 |
apply clarify
|
|
137 |
apply(erule transition_cases,simp_all)
|
|
138 |
apply(force dest:nth_mem simp add:All_None_def)
|
|
139 |
done
|
|
140 |
|
|
141 |
lemma Parallel_AllNone: "All_None Ts \<Longrightarrow> (SEM (Parallel Ts) X) = X"
|
|
142 |
apply (unfold SEM_def sem_def)
|
|
143 |
apply auto
|
|
144 |
apply(drule rtrancl_imp_UN_rel_pow)
|
|
145 |
apply clarify
|
|
146 |
apply(drule Parallel_AllNone_lemma)
|
|
147 |
apply auto
|
|
148 |
done
|
|
149 |
|
|
150 |
lemma Parallel_empty: "Ts=[] \<Longrightarrow> (SEM (Parallel Ts) X) = X"
|
|
151 |
apply(rule Parallel_AllNone)
|
|
152 |
apply(simp add:All_None_def)
|
|
153 |
done
|
|
154 |
|
|
155 |
text {* Set of lemmas from Apt and Olderog "Verification of sequential
|
|
156 |
and concurrent programs", page 63. *}
|
|
157 |
|
|
158 |
lemma L3_5i: "X\<subseteq>Y \<Longrightarrow> SEM c X \<subseteq> SEM c Y"
|
|
159 |
apply (unfold SEM_def)
|
|
160 |
apply force
|
|
161 |
done
|
|
162 |
|
|
163 |
lemma L3_5ii_lemma1:
|
|
164 |
"\<lbrakk> (c1, s1) -P*\<rightarrow> (Parallel Ts, s2); All_None Ts;
|
|
165 |
(c2, s2) -P*\<rightarrow> (Parallel Ss, s3); All_None Ss \<rbrakk>
|
|
166 |
\<Longrightarrow> (Seq c1 c2, s1) -P*\<rightarrow> (Parallel Ss, s3)"
|
|
167 |
apply(erule converse_rtrancl_induct2)
|
|
168 |
apply(force intro:converse_rtrancl_into_rtrancl)+
|
|
169 |
done
|
|
170 |
|
|
171 |
lemma L3_5ii_lemma2 [rule_format (no_asm)]:
|
|
172 |
"\<forall>c1 c2 s t. (Seq c1 c2, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow>
|
|
173 |
(All_None Ts) \<longrightarrow> (\<exists>y m Rs. (c1,s) -P*\<rightarrow> (Parallel Rs, y) \<and>
|
|
174 |
(All_None Rs) \<and> (c2, y) -Pm\<rightarrow> (Parallel Ts, t) \<and> m \<le> n)"
|
|
175 |
apply(induct "n")
|
|
176 |
apply(force)
|
|
177 |
apply(safe dest!: rel_pow_Suc_D2)
|
|
178 |
apply(erule transition_cases,simp_all)
|
|
179 |
apply (fast intro!: le_SucI)
|
|
180 |
apply (fast intro!: le_SucI elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl)
|
|
181 |
done
|
|
182 |
|
|
183 |
lemma L3_5ii_lemma3:
|
|
184 |
"\<lbrakk>(Seq c1 c2,s) -P*\<rightarrow> (Parallel Ts,t); All_None Ts\<rbrakk> \<Longrightarrow>
|
|
185 |
(\<exists>y Rs. (c1,s) -P*\<rightarrow> (Parallel Rs,y) \<and> All_None Rs
|
|
186 |
\<and> (c2,y) -P*\<rightarrow> (Parallel Ts,t))"
|
|
187 |
apply(drule rtrancl_imp_UN_rel_pow)
|
|
188 |
apply(fast dest: L3_5ii_lemma2 rel_pow_imp_rtrancl)
|
|
189 |
done
|
|
190 |
|
|
191 |
lemma L3_5ii: "SEM (Seq c1 c2) X = SEM c2 (SEM c1 X)"
|
|
192 |
apply (unfold SEM_def sem_def)
|
|
193 |
apply auto
|
|
194 |
apply(fast dest: L3_5ii_lemma3)
|
|
195 |
apply(fast elim: L3_5ii_lemma1)
|
|
196 |
done
|
|
197 |
|
|
198 |
lemma L3_5iii: "SEM (Seq (Seq c1 c2) c3) X = SEM (Seq c1 (Seq c2 c3)) X"
|
|
199 |
apply (simp (no_asm) add: L3_5ii)
|
|
200 |
done
|
|
201 |
|
|
202 |
lemma L3_5iv:
|
|
203 |
"SEM (Cond b c1 c2) X = (SEM c1 (X \<inter> b)) Un (SEM c2 (X \<inter> (-b)))"
|
|
204 |
apply (unfold SEM_def sem_def)
|
|
205 |
apply auto
|
|
206 |
apply(erule converse_rtranclE)
|
|
207 |
prefer 2
|
|
208 |
apply (erule transition_cases,simp_all)
|
|
209 |
apply(fast intro: converse_rtrancl_into_rtrancl elim: transition_cases)+
|
|
210 |
done
|
|
211 |
|
|
212 |
|
|
213 |
lemma L3_5v_lemma1[rule_format]:
|
|
214 |
"(S,s) -Pn\<rightarrow> (T,t) \<longrightarrow> S=\<Omega> \<longrightarrow> (\<not>(\<exists>Rs. T=(Parallel Rs) \<and> All_None Rs))"
|
|
215 |
apply (unfold UNIV_def)
|
|
216 |
apply(rule nat_less_induct)
|
|
217 |
apply safe
|
|
218 |
apply(erule rel_pow_E2)
|
|
219 |
apply simp_all
|
|
220 |
apply(erule transition_cases)
|
|
221 |
apply simp_all
|
|
222 |
apply(erule rel_pow_E2)
|
|
223 |
apply(simp add: Id_def)
|
|
224 |
apply(erule transition_cases,simp_all)
|
|
225 |
apply clarify
|
|
226 |
apply(erule transition_cases,simp_all)
|
|
227 |
apply(erule rel_pow_E2,simp)
|
|
228 |
apply clarify
|
|
229 |
apply(erule transition_cases)
|
|
230 |
apply simp+
|
|
231 |
apply clarify
|
|
232 |
apply(erule transition_cases)
|
|
233 |
apply simp_all
|
|
234 |
done
|
|
235 |
|
|
236 |
lemma L3_5v_lemma2: "\<lbrakk>(\<Omega>, s) -P*\<rightarrow> (Parallel Ts, t); All_None Ts \<rbrakk> \<Longrightarrow> False"
|
|
237 |
apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1)
|
|
238 |
done
|
|
239 |
|
|
240 |
lemma L3_5v_lemma3: "SEM (\<Omega>) S = {}"
|
|
241 |
apply (unfold SEM_def sem_def)
|
|
242 |
apply(fast dest: L3_5v_lemma2)
|
|
243 |
done
|
|
244 |
|
|
245 |
lemma L3_5v_lemma4 [rule_format]:
|
|
246 |
"\<forall>s. (While b i c, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>
|
|
247 |
(\<exists>k. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t))"
|
|
248 |
apply(rule nat_less_induct)
|
|
249 |
apply safe
|
|
250 |
apply(erule rel_pow_E2)
|
|
251 |
apply safe
|
|
252 |
apply(erule transition_cases,simp_all)
|
|
253 |
apply (rule_tac x = "1" in exI)
|
|
254 |
apply(force dest: Parallel_empty_lemma intro: converse_rtrancl_into_rtrancl simp add: Id_def)
|
|
255 |
apply safe
|
|
256 |
apply(drule L3_5ii_lemma2)
|
|
257 |
apply safe
|
|
258 |
apply(drule le_imp_less_Suc)
|
|
259 |
apply (erule allE , erule impE,assumption)
|
|
260 |
apply (erule allE , erule impE, assumption)
|
|
261 |
apply safe
|
|
262 |
apply (rule_tac x = "k+1" in exI)
|
|
263 |
apply(simp (no_asm))
|
|
264 |
apply(rule converse_rtrancl_into_rtrancl)
|
|
265 |
apply fast
|
|
266 |
apply(fast elim: L3_5ii_lemma1)
|
|
267 |
done
|
|
268 |
|
|
269 |
lemma L3_5v_lemma5 [rule_format]:
|
|
270 |
"\<forall>s. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>
|
|
271 |
(While b i c, s) -P*\<rightarrow> (Parallel Ts,t)"
|
|
272 |
apply(induct "k")
|
|
273 |
apply(force dest: L3_5v_lemma2)
|
|
274 |
apply safe
|
|
275 |
apply(erule converse_rtranclE)
|
|
276 |
apply simp_all
|
|
277 |
apply(erule transition_cases,simp_all)
|
|
278 |
apply(rule converse_rtrancl_into_rtrancl)
|
|
279 |
apply(fast)
|
|
280 |
apply(fast elim!: L3_5ii_lemma1 dest: L3_5ii_lemma3)
|
|
281 |
apply(drule rtrancl_imp_UN_rel_pow)
|
|
282 |
apply clarify
|
|
283 |
apply(erule rel_pow_E2)
|
|
284 |
apply simp_all
|
|
285 |
apply(erule transition_cases,simp_all)
|
|
286 |
apply(fast dest: Parallel_empty_lemma)
|
|
287 |
done
|
|
288 |
|
|
289 |
lemma L3_5v: "SEM (While b i c) = (\<lambda>x. (\<Union>k. SEM (fwhile b c k) x))"
|
|
290 |
apply(rule ext)
|
|
291 |
apply (simp add: SEM_def sem_def)
|
|
292 |
apply safe
|
|
293 |
apply(drule rtrancl_imp_UN_rel_pow,simp)
|
|
294 |
apply clarify
|
|
295 |
apply(fast dest:L3_5v_lemma4)
|
|
296 |
apply(fast intro: L3_5v_lemma5)
|
|
297 |
done
|
|
298 |
|
|
299 |
section {* Validity of Correctness Formulas *}
|
|
300 |
|
|
301 |
constdefs
|
|
302 |
com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50)
|
|
303 |
"\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
|
|
304 |
|
|
305 |
ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45)
|
|
306 |
"\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
|
|
307 |
|
|
308 |
end |