author | blanchet |
Tue, 17 Aug 2010 18:18:14 +0200 | |
changeset 38497 | 78c4988831d1 |
parent 35416 | d8d7d1b785af |
child 39246 | 9e58f0499f57 |
permissions | -rw-r--r-- |
13020 | 1 |
|
2 |
header {* \section{The Proof System} *} |
|
3 |
||
16417 | 4 |
theory OG_Hoare imports OG_Tran begin |
13020 | 5 |
|
6 |
consts assertions :: "'a ann_com \<Rightarrow> ('a assn) set" |
|
7 |
primrec |
|
8 |
"assertions (AnnBasic r f) = {r}" |
|
9 |
"assertions (AnnSeq c1 c2) = assertions c1 \<union> assertions c2" |
|
10 |
"assertions (AnnCond1 r b c1 c2) = {r} \<union> assertions c1 \<union> assertions c2" |
|
11 |
"assertions (AnnCond2 r b c) = {r} \<union> assertions c" |
|
12 |
"assertions (AnnWhile r b i c) = {r, i} \<union> assertions c" |
|
13 |
"assertions (AnnAwait r b c) = {r}" |
|
14 |
||
15 |
consts atomics :: "'a ann_com \<Rightarrow> ('a assn \<times> 'a com) set" |
|
16 |
primrec |
|
17 |
"atomics (AnnBasic r f) = {(r, Basic f)}" |
|
18 |
"atomics (AnnSeq c1 c2) = atomics c1 \<union> atomics c2" |
|
19 |
"atomics (AnnCond1 r b c1 c2) = atomics c1 \<union> atomics c2" |
|
20 |
"atomics (AnnCond2 r b c) = atomics c" |
|
21 |
"atomics (AnnWhile r b i c) = atomics c" |
|
22 |
"atomics (AnnAwait r b c) = {(r \<inter> b, c)}" |
|
23 |
||
24 |
consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" |
|
25 |
primrec "com (c, q) = c" |
|
26 |
||
27 |
consts post :: "'a ann_triple_op \<Rightarrow> 'a assn" |
|
28 |
primrec "post (c, q) = q" |
|
29 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
30 |
definition interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool" where |
13020 | 31 |
"interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or> |
32 |
(\<forall>(r,a) \<in> atomics (the co'). \<parallel>= (q \<inter> r) a q \<and> |
|
33 |
(co = None \<or> (\<forall>p \<in> assertions (the co). \<parallel>= (p \<inter> r) a p)))" |
|
34 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
35 |
definition interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" where |
13020 | 36 |
"interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow> |
37 |
interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " |
|
38 |
||
23746 | 39 |
inductive |
40 |
oghoare :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>- _//_//_)" [90,55,90] 50) |
|
41 |
and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(2\<turnstile> _// _)" [60,90] 45) |
|
42 |
where |
|
13020 | 43 |
AnnBasic: "r \<subseteq> {s. f s \<in> q} \<Longrightarrow> \<turnstile> (AnnBasic r f) q" |
44 |
||
23746 | 45 |
| AnnSeq: "\<lbrakk> \<turnstile> c0 pre c1; \<turnstile> c1 q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnSeq c0 c1) q" |
13020 | 46 |
|
23746 | 47 |
| AnnCond1: "\<lbrakk> r \<inter> b \<subseteq> pre c1; \<turnstile> c1 q; r \<inter> -b \<subseteq> pre c2; \<turnstile> c2 q\<rbrakk> |
13020 | 48 |
\<Longrightarrow> \<turnstile> (AnnCond1 r b c1 c2) q" |
23746 | 49 |
| AnnCond2: "\<lbrakk> r \<inter> b \<subseteq> pre c; \<turnstile> c q; r \<inter> -b \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnCond2 r b c) q" |
13020 | 50 |
|
23746 | 51 |
| AnnWhile: "\<lbrakk> r \<subseteq> i; i \<inter> b \<subseteq> pre c; \<turnstile> c i; i \<inter> -b \<subseteq> q \<rbrakk> |
13020 | 52 |
\<Longrightarrow> \<turnstile> (AnnWhile r b i c) q" |
53 |
||
23746 | 54 |
| AnnAwait: "\<lbrakk> atom_com c; \<parallel>- (r \<inter> b) c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r b c) q" |
13020 | 55 |
|
23746 | 56 |
| AnnConseq: "\<lbrakk>\<turnstile> c q; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<turnstile> c q'" |
13020 | 57 |
|
58 |
||
23746 | 59 |
| Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32621
diff
changeset
|
60 |
\<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) |
13020 | 61 |
Parallel Ts |
62 |
(\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))" |
|
63 |
||
23746 | 64 |
| Basic: "\<parallel>- {s. f s \<in>q} (Basic f) q" |
13020 | 65 |
|
23746 | 66 |
| Seq: "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q " |
13020 | 67 |
|
23746 | 68 |
| Cond: "\<lbrakk> \<parallel>- (p \<inter> b) c1 q; \<parallel>- (p \<inter> -b) c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Cond b c1 c2) q" |
13020 | 69 |
|
23746 | 70 |
| While: "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)" |
13020 | 71 |
|
23746 | 72 |
| Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- p' c q'" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32621
diff
changeset
|
73 |
|
13020 | 74 |
section {* Soundness *} |
75 |
(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE |
|
76 |
parts of conditional expressions (if P then x else y) are no longer |
|
77 |
simplified. (This allows the simplifier to unfold recursive |
|
78 |
functional programs.) To restore the old behaviour, we declare |
|
79 |
@{text "lemmas [cong del] = if_weak_cong"}. *) |
|
80 |
||
81 |
lemmas [cong del] = if_weak_cong |
|
82 |
||
83 |
lemmas ann_hoare_induct = oghoare_ann_hoare.induct [THEN conjunct2] |
|
84 |
lemmas oghoare_induct = oghoare_ann_hoare.induct [THEN conjunct1] |
|
85 |
||
86 |
lemmas AnnBasic = oghoare_ann_hoare.AnnBasic |
|
87 |
lemmas AnnSeq = oghoare_ann_hoare.AnnSeq |
|
88 |
lemmas AnnCond1 = oghoare_ann_hoare.AnnCond1 |
|
89 |
lemmas AnnCond2 = oghoare_ann_hoare.AnnCond2 |
|
90 |
lemmas AnnWhile = oghoare_ann_hoare.AnnWhile |
|
91 |
lemmas AnnAwait = oghoare_ann_hoare.AnnAwait |
|
92 |
lemmas AnnConseq = oghoare_ann_hoare.AnnConseq |
|
93 |
||
94 |
lemmas Parallel = oghoare_ann_hoare.Parallel |
|
95 |
lemmas Basic = oghoare_ann_hoare.Basic |
|
96 |
lemmas Seq = oghoare_ann_hoare.Seq |
|
97 |
lemmas Cond = oghoare_ann_hoare.Cond |
|
98 |
lemmas While = oghoare_ann_hoare.While |
|
99 |
lemmas Conseq = oghoare_ann_hoare.Conseq |
|
100 |
||
101 |
subsection {* Soundness of the System for Atomic Programs *} |
|
102 |
||
103 |
lemma Basic_ntran [rule_format]: |
|
104 |
"(Basic f, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow> t = f s" |
|
105 |
apply(induct "n") |
|
106 |
apply(simp (no_asm)) |
|
107 |
apply(fast dest: rel_pow_Suc_D2 Parallel_empty_lemma elim: transition_cases) |
|
108 |
done |
|
109 |
||
110 |
lemma SEM_fwhile: "SEM S (p \<inter> b) \<subseteq> p \<Longrightarrow> SEM (fwhile b S k) p \<subseteq> (p \<inter> -b)" |
|
111 |
apply (induct "k") |
|
112 |
apply(simp (no_asm) add: L3_5v_lemma3) |
|
113 |
apply(simp (no_asm) add: L3_5iv L3_5ii Parallel_empty) |
|
15102 | 114 |
apply(rule conjI) |
115 |
apply (blast dest: L3_5i) |
|
13020 | 116 |
apply(simp add: SEM_def sem_def id_def) |
15102 | 117 |
apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow) |
13020 | 118 |
done |
119 |
||
15102 | 120 |
lemma atom_hoare_sound [rule_format]: |
13020 | 121 |
" \<parallel>- p c q \<longrightarrow> atom_com(c) \<longrightarrow> \<parallel>= p c q" |
122 |
apply (unfold com_validity_def) |
|
123 |
apply(rule oghoare_induct) |
|
124 |
apply simp_all |
|
125 |
--{*Basic*} |
|
126 |
apply(simp add: SEM_def sem_def) |
|
127 |
apply(fast dest: rtrancl_imp_UN_rel_pow Basic_ntran) |
|
128 |
--{* Seq *} |
|
129 |
apply(rule impI) |
|
130 |
apply(rule subset_trans) |
|
131 |
prefer 2 apply simp |
|
132 |
apply(simp add: L3_5ii L3_5i) |
|
133 |
--{* Cond *} |
|
134 |
apply(simp add: L3_5iv) |
|
135 |
--{* While *} |
|
15102 | 136 |
apply (force simp add: L3_5v dest: SEM_fwhile) |
13020 | 137 |
--{* Conseq *} |
15102 | 138 |
apply(force simp add: SEM_def sem_def) |
13020 | 139 |
done |
140 |
||
141 |
subsection {* Soundness of the System for Component Programs *} |
|
142 |
||
143 |
inductive_cases ann_transition_cases: |
|
23746 | 144 |
"(None,s) -1\<rightarrow> (c', s')" |
145 |
"(Some (AnnBasic r f),s) -1\<rightarrow> (c', s')" |
|
146 |
"(Some (AnnSeq c1 c2), s) -1\<rightarrow> (c', s')" |
|
147 |
"(Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (c', s')" |
|
148 |
"(Some (AnnCond2 r b c), s) -1\<rightarrow> (c', s')" |
|
149 |
"(Some (AnnWhile r b I c), s) -1\<rightarrow> (c', s')" |
|
150 |
"(Some (AnnAwait r b c),s) -1\<rightarrow> (c', s')" |
|
13020 | 151 |
|
152 |
text {* Strong Soundness for Component Programs:*} |
|
153 |
||
26811
067cceb36e26
Rephrased proof of ann_hoare_case_analysis, to avoid problems with HO unification
berghofe
parents:
23746
diff
changeset
|
154 |
lemma ann_hoare_case_analysis [rule_format]: |
067cceb36e26
Rephrased proof of ann_hoare_case_analysis, to avoid problems with HO unification
berghofe
parents:
23746
diff
changeset
|
155 |
defines I: "I \<equiv> \<lambda>C q'. |
13020 | 156 |
((\<forall>r f. C = AnnBasic r f \<longrightarrow> (\<exists>q. r \<subseteq> {s. f s \<in> q} \<and> q \<subseteq> q')) \<and> |
157 |
(\<forall>c0 c1. C = AnnSeq c0 c1 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<turnstile> c0 pre c1 \<and> \<turnstile> c1 q)) \<and> |
|
158 |
(\<forall>r b c1 c2. C = AnnCond1 r b c1 c2 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> |
|
159 |
r \<inter> b \<subseteq> pre c1 \<and> \<turnstile> c1 q \<and> r \<inter> -b \<subseteq> pre c2 \<and> \<turnstile> c2 q)) \<and> |
|
160 |
(\<forall>r b c. C = AnnCond2 r b c \<longrightarrow> |
|
161 |
(\<exists>q. q \<subseteq> q' \<and> r \<inter> b \<subseteq> pre c \<and> \<turnstile> c q \<and> r \<inter> -b \<subseteq> q)) \<and> |
|
162 |
(\<forall>r i b c. C = AnnWhile r b i c \<longrightarrow> |
|
163 |
(\<exists>q. q \<subseteq> q' \<and> r \<subseteq> i \<and> i \<inter> b \<subseteq> pre c \<and> \<turnstile> c i \<and> i \<inter> -b \<subseteq> q)) \<and> |
|
164 |
(\<forall>r b c. C = AnnAwait r b c \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<parallel>- (r \<inter> b) c q)))" |
|
26811
067cceb36e26
Rephrased proof of ann_hoare_case_analysis, to avoid problems with HO unification
berghofe
parents:
23746
diff
changeset
|
165 |
shows "\<turnstile> C q' \<longrightarrow> I C q'" |
13020 | 166 |
apply(rule ann_hoare_induct) |
26811
067cceb36e26
Rephrased proof of ann_hoare_case_analysis, to avoid problems with HO unification
berghofe
parents:
23746
diff
changeset
|
167 |
apply (simp_all add: I) |
13020 | 168 |
apply(rule_tac x=q in exI,simp)+ |
169 |
apply(rule conjI,clarify,simp,clarify,rule_tac x=qa in exI,fast)+ |
|
170 |
apply(clarify,simp,clarify,rule_tac x=qa in exI,fast) |
|
171 |
done |
|
172 |
||
23746 | 173 |
lemma Help: "(transition \<inter> {(x,y). True}) = (transition)" |
13020 | 174 |
apply force |
175 |
done |
|
176 |
||
177 |
lemma Strong_Soundness_aux_aux [rule_format]: |
|
178 |
"(co, s) -1\<rightarrow> (co', t) \<longrightarrow> (\<forall>c. co = Some c \<longrightarrow> s\<in> pre c \<longrightarrow> |
|
179 |
(\<forall>q. \<turnstile> c q \<longrightarrow> (if co' = None then t\<in>q else t \<in> pre(the co') \<and> \<turnstile> (the co') q )))" |
|
180 |
apply(rule ann_transition_transition.induct [THEN conjunct1]) |
|
181 |
apply simp_all |
|
182 |
--{* Basic *} |
|
183 |
apply clarify |
|
184 |
apply(frule ann_hoare_case_analysis) |
|
185 |
apply force |
|
186 |
--{* Seq *} |
|
187 |
apply clarify |
|
188 |
apply(frule ann_hoare_case_analysis,simp) |
|
189 |
apply(fast intro: AnnConseq) |
|
190 |
apply clarify |
|
191 |
apply(frule ann_hoare_case_analysis,simp) |
|
192 |
apply clarify |
|
193 |
apply(rule conjI) |
|
194 |
apply force |
|
195 |
apply(rule AnnSeq,simp) |
|
196 |
apply(fast intro: AnnConseq) |
|
197 |
--{* Cond1 *} |
|
198 |
apply clarify |
|
199 |
apply(frule ann_hoare_case_analysis,simp) |
|
200 |
apply(fast intro: AnnConseq) |
|
201 |
apply clarify |
|
202 |
apply(frule ann_hoare_case_analysis,simp) |
|
203 |
apply(fast intro: AnnConseq) |
|
204 |
--{* Cond2 *} |
|
205 |
apply clarify |
|
206 |
apply(frule ann_hoare_case_analysis,simp) |
|
207 |
apply(fast intro: AnnConseq) |
|
208 |
apply clarify |
|
209 |
apply(frule ann_hoare_case_analysis,simp) |
|
210 |
apply(fast intro: AnnConseq) |
|
211 |
--{* While *} |
|
212 |
apply clarify |
|
213 |
apply(frule ann_hoare_case_analysis,simp) |
|
214 |
apply force |
|
215 |
apply clarify |
|
216 |
apply(frule ann_hoare_case_analysis,simp) |
|
217 |
apply auto |
|
218 |
apply(rule AnnSeq) |
|
219 |
apply simp |
|
220 |
apply(rule AnnWhile) |
|
221 |
apply simp_all |
|
222 |
--{* Await *} |
|
223 |
apply(frule ann_hoare_case_analysis,simp) |
|
224 |
apply clarify |
|
225 |
apply(drule atom_hoare_sound) |
|
226 |
apply simp |
|
227 |
apply(simp add: com_validity_def SEM_def sem_def) |
|
228 |
apply(simp add: Help All_None_def) |
|
229 |
apply force |
|
230 |
done |
|
231 |
||
232 |
lemma Strong_Soundness_aux: "\<lbrakk> (Some c, s) -*\<rightarrow> (co, t); s \<in> pre c; \<turnstile> c q \<rbrakk> |
|
233 |
\<Longrightarrow> if co = None then t \<in> q else t \<in> pre (the co) \<and> \<turnstile> (the co) q" |
|
234 |
apply(erule rtrancl_induct2) |
|
235 |
apply simp |
|
236 |
apply(case_tac "a") |
|
237 |
apply(fast elim: ann_transition_cases) |
|
238 |
apply(erule Strong_Soundness_aux_aux) |
|
239 |
apply simp |
|
240 |
apply simp_all |
|
241 |
done |
|
242 |
||
243 |
lemma Strong_Soundness: "\<lbrakk> (Some c, s)-*\<rightarrow>(co, t); s \<in> pre c; \<turnstile> c q \<rbrakk> |
|
244 |
\<Longrightarrow> if co = None then t\<in>q else t \<in> pre (the co)" |
|
245 |
apply(force dest:Strong_Soundness_aux) |
|
246 |
done |
|
247 |
||
248 |
lemma ann_hoare_sound: "\<turnstile> c q \<Longrightarrow> \<Turnstile> c q" |
|
249 |
apply (unfold ann_com_validity_def ann_SEM_def ann_sem_def) |
|
250 |
apply clarify |
|
251 |
apply(drule Strong_Soundness) |
|
252 |
apply simp_all |
|
253 |
done |
|
254 |
||
255 |
subsection {* Soundness of the System for Parallel Programs *} |
|
256 |
||
257 |
lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1\<rightarrow> (R', t) \<Longrightarrow> |
|
258 |
(\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and> |
|
259 |
(\<forall>i. i<length Ts \<longrightarrow> post(Rs ! i) = post(Ts ! i)))" |
|
260 |
apply(erule transition_cases) |
|
261 |
apply simp |
|
262 |
apply clarify |
|
263 |
apply(case_tac "i=ia") |
|
264 |
apply simp+ |
|
265 |
done |
|
266 |
||
267 |
lemma Parallel_length_post_PStar: "(Parallel Ts,s) -P*\<rightarrow> (R',t) \<Longrightarrow> |
|
268 |
(\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and> |
|
269 |
(\<forall>i. i<length Ts \<longrightarrow> post(Ts ! i) = post(Rs ! i)))" |
|
270 |
apply(erule rtrancl_induct2) |
|
271 |
apply(simp_all) |
|
272 |
apply clarify |
|
273 |
apply simp |
|
274 |
apply(drule Parallel_length_post_P1) |
|
275 |
apply auto |
|
276 |
done |
|
277 |
||
278 |
lemma assertions_lemma: "pre c \<in> assertions c" |
|
279 |
apply(rule ann_com_com.induct [THEN conjunct1]) |
|
280 |
apply auto |
|
281 |
done |
|
282 |
||
283 |
lemma interfree_aux1 [rule_format]: |
|
284 |
"(c,s) -1\<rightarrow> (r,t) \<longrightarrow> (interfree_aux(c1, q1, c) \<longrightarrow> interfree_aux(c1, q1, r))" |
|
285 |
apply (rule ann_transition_transition.induct [THEN conjunct1]) |
|
286 |
apply(safe) |
|
287 |
prefer 13 |
|
288 |
apply (rule TrueI) |
|
289 |
apply (simp_all add:interfree_aux_def) |
|
290 |
apply force+ |
|
291 |
done |
|
292 |
||
293 |
lemma interfree_aux2 [rule_format]: |
|
294 |
"(c,s) -1\<rightarrow> (r,t) \<longrightarrow> (interfree_aux(c, q, a) \<longrightarrow> interfree_aux(r, q, a) )" |
|
295 |
apply (rule ann_transition_transition.induct [THEN conjunct1]) |
|
296 |
apply(force simp add:interfree_aux_def)+ |
|
297 |
done |
|
298 |
||
299 |
lemma interfree_lemma: "\<lbrakk> (Some c, s) -1\<rightarrow> (r, t);interfree Ts ; i<length Ts; |
|
300 |
Ts!i = (Some c, q) \<rbrakk> \<Longrightarrow> interfree (Ts[i:= (r, q)])" |
|
301 |
apply(simp add: interfree_def) |
|
302 |
apply clarify |
|
303 |
apply(case_tac "i=j") |
|
304 |
apply(drule_tac t = "ia" in not_sym) |
|
305 |
apply simp_all |
|
306 |
apply(force elim: interfree_aux1) |
|
307 |
apply(force elim: interfree_aux2 simp add:nth_list_update) |
|
308 |
done |
|
309 |
||
310 |
text {* Strong Soundness Theorem for Parallel Programs:*} |
|
311 |
||
312 |
lemma Parallel_Strong_Soundness_Seq_aux: |
|
313 |
"\<lbrakk>interfree Ts; i<length Ts; com(Ts ! i) = Some(AnnSeq c0 c1) \<rbrakk> |
|
314 |
\<Longrightarrow> interfree (Ts[i:=(Some c0, pre c1)])" |
|
315 |
apply(simp add: interfree_def) |
|
316 |
apply clarify |
|
317 |
apply(case_tac "i=j") |
|
318 |
apply(force simp add: nth_list_update interfree_aux_def) |
|
319 |
apply(case_tac "i=ia") |
|
320 |
apply(erule_tac x=ia in allE) |
|
321 |
apply(force simp add:interfree_aux_def assertions_lemma) |
|
322 |
apply simp |
|
323 |
done |
|
324 |
||
325 |
lemma Parallel_Strong_Soundness_Seq [rule_format (no_asm)]: |
|
326 |
"\<lbrakk> \<forall>i<length Ts. (if com(Ts!i) = None then b \<in> post(Ts!i) |
|
327 |
else b \<in> pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i)); |
|
328 |
com(Ts ! i) = Some(AnnSeq c0 c1); i<length Ts; interfree Ts \<rbrakk> \<Longrightarrow> |
|
329 |
(\<forall>ia<length Ts. (if com(Ts[i:=(Some c0, pre c1)]! ia) = None |
|
330 |
then b \<in> post(Ts[i:=(Some c0, pre c1)]! ia) |
|
331 |
else b \<in> pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \<and> |
|
332 |
\<turnstile> the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) |
|
333 |
\<and> interfree (Ts[i:= (Some c0, pre c1)])" |
|
334 |
apply(rule conjI) |
|
335 |
apply safe |
|
336 |
apply(case_tac "i=ia") |
|
337 |
apply simp |
|
338 |
apply(force dest: ann_hoare_case_analysis) |
|
339 |
apply simp |
|
340 |
apply(fast elim: Parallel_Strong_Soundness_Seq_aux) |
|
341 |
done |
|
342 |
||
343 |
lemma Parallel_Strong_Soundness_aux_aux [rule_format]: |
|
344 |
"(Some c, b) -1\<rightarrow> (co, t) \<longrightarrow> |
|
345 |
(\<forall>Ts. i<length Ts \<longrightarrow> com(Ts ! i) = Some c \<longrightarrow> |
|
346 |
(\<forall>i<length Ts. (if com(Ts ! i) = None then b\<in>post(Ts!i) |
|
347 |
else b\<in>pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i))) \<longrightarrow> |
|
348 |
interfree Ts \<longrightarrow> |
|
349 |
(\<forall>j. j<length Ts \<and> i\<noteq>j \<longrightarrow> (if com(Ts!j) = None then t\<in>post(Ts!j) |
|
350 |
else t\<in>pre(the(com(Ts!j))) \<and> \<turnstile> the(com(Ts!j)) post(Ts!j))) )" |
|
351 |
apply(rule ann_transition_transition.induct [THEN conjunct1]) |
|
352 |
apply safe |
|
353 |
prefer 11 |
|
354 |
apply(rule TrueI) |
|
355 |
apply simp_all |
|
356 |
--{* Basic *} |
|
357 |
apply(erule_tac x = "i" in all_dupE, erule (1) notE impE) |
|
358 |
apply(erule_tac x = "j" in allE , erule (1) notE impE) |
|
359 |
apply(simp add: interfree_def) |
|
360 |
apply(erule_tac x = "j" in allE,simp) |
|
361 |
apply(erule_tac x = "i" in allE,simp) |
|
362 |
apply(drule_tac t = "i" in not_sym) |
|
363 |
apply(case_tac "com(Ts ! j)=None") |
|
364 |
apply(force intro: converse_rtrancl_into_rtrancl |
|
365 |
simp add: interfree_aux_def com_validity_def SEM_def sem_def All_None_def) |
|
366 |
apply(simp add:interfree_aux_def) |
|
367 |
apply clarify |
|
368 |
apply simp |
|
369 |
apply(erule_tac x="pre y" in ballE) |
|
370 |
apply(force intro: converse_rtrancl_into_rtrancl |
|
371 |
simp add: com_validity_def SEM_def sem_def All_None_def) |
|
372 |
apply(simp add:assertions_lemma) |
|
373 |
--{* Seqs *} |
|
374 |
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) |
|
375 |
apply(drule Parallel_Strong_Soundness_Seq,simp+) |
|
376 |
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) |
|
377 |
apply(drule Parallel_Strong_Soundness_Seq,simp+) |
|
378 |
--{* Await *} |
|
379 |
apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE) |
|
380 |
apply(erule_tac x = "j" in allE , erule (1) notE impE) |
|
381 |
apply(simp add: interfree_def) |
|
382 |
apply(erule_tac x = "j" in allE,simp) |
|
383 |
apply(erule_tac x = "i" in allE,simp) |
|
384 |
apply(drule_tac t = "i" in not_sym) |
|
385 |
apply(case_tac "com(Ts ! j)=None") |
|
386 |
apply(force intro: converse_rtrancl_into_rtrancl simp add: interfree_aux_def |
|
387 |
com_validity_def SEM_def sem_def All_None_def Help) |
|
388 |
apply(simp add:interfree_aux_def) |
|
389 |
apply clarify |
|
390 |
apply simp |
|
391 |
apply(erule_tac x="pre y" in ballE) |
|
392 |
apply(force intro: converse_rtrancl_into_rtrancl |
|
393 |
simp add: com_validity_def SEM_def sem_def All_None_def Help) |
|
394 |
apply(simp add:assertions_lemma) |
|
395 |
done |
|
396 |
||
397 |
lemma Parallel_Strong_Soundness_aux [rule_format]: |
|
398 |
"\<lbrakk>(Ts',s) -P*\<rightarrow> (Rs',t); Ts' = (Parallel Ts); interfree Ts; |
|
399 |
\<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. (Ts ! i) = (Some c, q) \<and> s\<in>(pre c) \<and> \<turnstile> c q ) \<rbrakk> \<Longrightarrow> |
|
400 |
\<forall>Rs. Rs' = (Parallel Rs) \<longrightarrow> (\<forall>j. j<length Rs \<longrightarrow> |
|
401 |
(if com(Rs ! j) = None then t\<in>post(Ts ! j) |
|
402 |
else t\<in>pre(the(com(Rs ! j))) \<and> \<turnstile> the(com(Rs ! j)) post(Ts ! j))) \<and> interfree Rs" |
|
403 |
apply(erule rtrancl_induct2) |
|
404 |
apply clarify |
|
405 |
--{* Base *} |
|
406 |
apply force |
|
407 |
--{* Induction step *} |
|
408 |
apply clarify |
|
409 |
apply(drule Parallel_length_post_PStar) |
|
410 |
apply clarify |
|
23746 | 411 |
apply (ind_cases "(Parallel Ts, s) -P1\<rightarrow> (Parallel Rs, t)" for Ts s Rs t) |
13020 | 412 |
apply(rule conjI) |
413 |
apply clarify |
|
414 |
apply(case_tac "i=j") |
|
415 |
apply(simp split del:split_if) |
|
416 |
apply(erule Strong_Soundness_aux_aux,simp+) |
|
417 |
apply force |
|
418 |
apply force |
|
419 |
apply(simp split del: split_if) |
|
420 |
apply(erule Parallel_Strong_Soundness_aux_aux) |
|
421 |
apply(simp_all add: split del:split_if) |
|
422 |
apply force |
|
423 |
apply(rule interfree_lemma) |
|
424 |
apply simp_all |
|
425 |
done |
|
426 |
||
427 |
lemma Parallel_Strong_Soundness: |
|
428 |
"\<lbrakk>(Parallel Ts, s) -P*\<rightarrow> (Parallel Rs, t); interfree Ts; j<length Rs; |
|
429 |
\<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. Ts ! i = (Some c, q) \<and> s\<in>pre c \<and> \<turnstile> c q) \<rbrakk> \<Longrightarrow> |
|
430 |
if com(Rs ! j) = None then t\<in>post(Ts ! j) else t\<in>pre (the(com(Rs ! j)))" |
|
431 |
apply(drule Parallel_Strong_Soundness_aux) |
|
432 |
apply simp+ |
|
433 |
done |
|
434 |
||
15102 | 435 |
lemma oghoare_sound [rule_format]: "\<parallel>- p c q \<longrightarrow> \<parallel>= p c q" |
13020 | 436 |
apply (unfold com_validity_def) |
437 |
apply(rule oghoare_induct) |
|
438 |
apply(rule TrueI)+ |
|
439 |
--{* Parallel *} |
|
440 |
apply(simp add: SEM_def sem_def) |
|
441 |
apply clarify |
|
442 |
apply(frule Parallel_length_post_PStar) |
|
443 |
apply clarify |
|
32139 | 444 |
apply(drule_tac j=xb in Parallel_Strong_Soundness) |
13020 | 445 |
apply clarify |
446 |
apply simp |
|
447 |
apply force |
|
448 |
apply simp |
|
449 |
apply(erule_tac V = "\<forall>i. ?P i" in thin_rl) |
|
450 |
apply(drule_tac s = "length Rs" in sym) |
|
451 |
apply(erule allE, erule impE, assumption) |
|
452 |
apply(force dest: nth_mem simp add: All_None_def) |
|
453 |
--{* Basic *} |
|
454 |
apply(simp add: SEM_def sem_def) |
|
455 |
apply(force dest: rtrancl_imp_UN_rel_pow Basic_ntran) |
|
456 |
--{* Seq *} |
|
457 |
apply(rule subset_trans) |
|
458 |
prefer 2 apply assumption |
|
459 |
apply(simp add: L3_5ii L3_5i) |
|
460 |
--{* Cond *} |
|
461 |
apply(simp add: L3_5iv) |
|
462 |
--{* While *} |
|
463 |
apply(simp add: L3_5v) |
|
15102 | 464 |
apply (blast dest: SEM_fwhile) |
13020 | 465 |
--{* Conseq *} |
15102 | 466 |
apply(auto simp add: SEM_def sem_def) |
13020 | 467 |
done |
468 |
||
469 |
end |