author | blanchet |
Tue, 17 Feb 2015 17:22:45 +0100 | |
changeset 59552 | ae50c9b82444 |
parent 59498 | 50b60f501b05 |
child 59763 | 56d2c357e6b5 |
permissions | -rw-r--r-- |
8177 | 1 |
(* Title: HOL/IMPP/Hoare.thy |
2 |
Author: David von Oheimb |
|
3 |
Copyright 1999 TUM |
|
4 |
*) |
|
5 |
||
58889 | 6 |
section {* Inductive definition of Hoare logic for partial correctness *} |
17477 | 7 |
|
8 |
theory Hoare |
|
9 |
imports Natural |
|
10 |
begin |
|
11 |
||
12 |
text {* |
|
13 |
Completeness is taken relative to completeness of the underlying logic. |
|
14 |
||
15 |
Two versions of completeness proof: nested single recursion |
|
16 |
vs. simultaneous recursion in call rule |
|
17 |
*} |
|
8177 | 18 |
|
42174 | 19 |
type_synonym 'a assn = "'a => state => bool" |
8177 | 20 |
translations |
35431 | 21 |
(type) "'a assn" <= (type) "'a => state => bool" |
8177 | 22 |
|
27364 | 23 |
definition |
24 |
state_not_singleton :: bool where |
|
25 |
"state_not_singleton = (\<exists>s t::state. s ~= t)" (* at least two elements *) |
|
8177 | 26 |
|
27364 | 27 |
definition |
28 |
peek_and :: "'a assn => (state => bool) => 'a assn" (infixr "&>" 35) where |
|
29 |
"peek_and P p = (%Z s. P Z s & p s)" |
|
8177 | 30 |
|
58310 | 31 |
datatype 'a triple = |
17477 | 32 |
triple "'a assn" com "'a assn" ("{(1_)}./ (_)/ .{(1_)}" [3,60,3] 58) |
33 |
||
27364 | 34 |
definition |
35 |
triple_valid :: "nat => 'a triple => bool" ( "|=_:_" [0 , 58] 57) where |
|
36 |
"|=n:t = (case t of {P}.c.{Q} => |
|
37 |
!Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s'))" |
|
38 |
abbreviation |
|
39 |
triples_valid :: "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57) where |
|
40 |
"||=n:G == Ball G (triple_valid n)" |
|
8177 | 41 |
|
27364 | 42 |
definition |
43 |
hoare_valids :: "'a triple set => 'a triple set => bool" ("_||=_" [58, 58] 57) where |
|
44 |
"G||=ts = (!n. ||=n:G --> ||=n:ts)" |
|
45 |
abbreviation |
|
46 |
hoare_valid :: "'a triple set => 'a triple => bool" ("_|=_" [58, 58] 57) where |
|
47 |
"G |=t == G||={t}" |
|
8177 | 48 |
|
49 |
(* Most General Triples *) |
|
27364 | 50 |
definition |
51 |
MGT :: "com => state triple" ("{=}._.{->}" [60] 58) where |
|
52 |
"{=}.c.{->} = {%Z s0. Z = s0}. c .{%Z s1. <c,Z> -c-> s1}" |
|
8177 | 53 |
|
23746 | 54 |
inductive |
27364 | 55 |
hoare_derivs :: "'a triple set => 'a triple set => bool" ("_||-_" [58, 58] 57) and |
56 |
hoare_deriv :: "'a triple set => 'a triple => bool" ("_|-_" [58, 58] 57) |
|
23746 | 57 |
where |
58 |
"G |-t == G||-{t}" |
|
17477 | 59 |
|
23746 | 60 |
| empty: "G||-{}" |
61 |
| insert: "[| G |-t; G||-ts |] |
|
17477 | 62 |
==> G||-insert t ts" |
8177 | 63 |
|
23746 | 64 |
| asm: "ts <= G ==> |
17477 | 65 |
G||-ts" (* {P}.BODY pn.{Q} instead of (general) t for SkipD_lemma *) |
8177 | 66 |
|
23746 | 67 |
| cut: "[| G'||-ts; G||-G' |] ==> G||-ts" (* for convenience and efficiency *) |
8177 | 68 |
|
23746 | 69 |
| weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts" |
8177 | 70 |
|
23746 | 71 |
| conseq: "!Z s. P Z s --> (? P' Q'. G|-{P'}.c.{Q'} & |
17477 | 72 |
(!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s')) |
73 |
==> G|-{P}.c.{Q}" |
|
8177 | 74 |
|
75 |
||
23746 | 76 |
| Skip: "G|-{P}. SKIP .{P}" |
8177 | 77 |
|
23746 | 78 |
| Ass: "G|-{%Z s. P Z (s[X::=a s])}. X:==a .{P}" |
8177 | 79 |
|
23746 | 80 |
| Local: "G|-{P}. c .{%Z s. Q Z (s[Loc X::=s'<X>])} |
17477 | 81 |
==> G|-{%Z s. s'=s & P Z (s[Loc X::=a s])}. LOCAL X:=a IN c .{Q}" |
8177 | 82 |
|
23746 | 83 |
| Comp: "[| G|-{P}.c.{Q}; |
17477 | 84 |
G|-{Q}.d.{R} |] |
85 |
==> G|-{P}. (c;;d) .{R}" |
|
8177 | 86 |
|
23746 | 87 |
| If: "[| G|-{P &> b }.c.{Q}; |
17477 | 88 |
G|-{P &> (Not o b)}.d.{Q} |] |
89 |
==> G|-{P}. IF b THEN c ELSE d .{Q}" |
|
8177 | 90 |
|
23746 | 91 |
| Loop: "G|-{P &> b}.c.{P} ==> |
17477 | 92 |
G|-{P}. WHILE b DO c .{P &> (Not o b)}" |
8177 | 93 |
|
94 |
(* |
|
17477 | 95 |
BodyN: "(insert ({P}. BODY pn .{Q}) G) |
96 |
|-{P}. the (body pn) .{Q} ==> |
|
97 |
G|-{P}. BODY pn .{Q}" |
|
8177 | 98 |
*) |
23746 | 99 |
| Body: "[| G Un (%p. {P p}. BODY p .{Q p})`Procs |
17477 | 100 |
||-(%p. {P p}. the (body p) .{Q p})`Procs |] |
101 |
==> G||-(%p. {P p}. BODY p .{Q p})`Procs" |
|
8177 | 102 |
|
23746 | 103 |
| Call: "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])} |
17477 | 104 |
==> G|-{%Z s. s'=s & P Z (setlocs s newlocs[Loc Arg::=a s])}. |
105 |
X:=CALL pn(a) .{Q}" |
|
106 |
||
19803 | 107 |
|
108 |
section {* Soundness and relative completeness of Hoare rules wrt operational semantics *} |
|
109 |
||
58254 | 110 |
lemma single_stateE: |
19803 | 111 |
"state_not_singleton ==> !t. (!s::state. s = t) --> False" |
112 |
apply (unfold state_not_singleton_def) |
|
113 |
apply clarify |
|
114 |
apply (case_tac "ta = t") |
|
115 |
apply blast |
|
116 |
apply (blast dest: not_sym) |
|
117 |
done |
|
118 |
||
119 |
declare peek_and_def [simp] |
|
120 |
||
121 |
||
122 |
subsection "validity" |
|
123 |
||
58254 | 124 |
lemma triple_valid_def2: |
19803 | 125 |
"|=n:{P}.c.{Q} = (!Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s'))" |
126 |
apply (unfold triple_valid_def) |
|
127 |
apply auto |
|
128 |
done |
|
129 |
||
130 |
lemma Body_triple_valid_0: "|=0:{P}. BODY pn .{Q}" |
|
131 |
apply (simp (no_asm) add: triple_valid_def2) |
|
132 |
apply clarsimp |
|
133 |
done |
|
134 |
||
135 |
(* only ==> direction required *) |
|
136 |
lemma Body_triple_valid_Suc: "|=n:{P}. the (body pn) .{Q} = |=Suc n:{P}. BODY pn .{Q}" |
|
137 |
apply (simp (no_asm) add: triple_valid_def2) |
|
138 |
apply force |
|
139 |
done |
|
140 |
||
141 |
lemma triple_valid_Suc [rule_format (no_asm)]: "|=Suc n:t --> |=n:t" |
|
142 |
apply (unfold triple_valid_def) |
|
143 |
apply (induct_tac t) |
|
144 |
apply simp |
|
145 |
apply (fast intro: evaln_Suc) |
|
146 |
done |
|
147 |
||
148 |
lemma triples_valid_Suc: "||=Suc n:ts ==> ||=n:ts" |
|
149 |
apply (fast intro: triple_valid_Suc) |
|
150 |
done |
|
151 |
||
152 |
||
153 |
subsection "derived rules" |
|
154 |
||
58254 | 155 |
lemma conseq12: "[| G|-{P'}.c.{Q'}; !Z s. P Z s --> |
156 |
(!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s') |] |
|
19803 | 157 |
==> G|-{P}.c.{Q}" |
158 |
apply (rule hoare_derivs.conseq) |
|
159 |
apply blast |
|
160 |
done |
|
161 |
||
162 |
lemma conseq1: "[| G|-{P'}.c.{Q}; !Z s. P Z s --> P' Z s |] ==> G|-{P}.c.{Q}" |
|
163 |
apply (erule conseq12) |
|
164 |
apply fast |
|
165 |
done |
|
166 |
||
167 |
lemma conseq2: "[| G|-{P}.c.{Q'}; !Z s. Q' Z s --> Q Z s |] ==> G|-{P}.c.{Q}" |
|
168 |
apply (erule conseq12) |
|
169 |
apply fast |
|
170 |
done |
|
171 |
||
58254 | 172 |
lemma Body1: "[| G Un (%p. {P p}. BODY p .{Q p})`Procs |
173 |
||- (%p. {P p}. the (body p) .{Q p})`Procs; |
|
19803 | 174 |
pn:Procs |] ==> G|-{P pn}. BODY pn .{Q pn}" |
175 |
apply (drule hoare_derivs.Body) |
|
176 |
apply (erule hoare_derivs.weaken) |
|
177 |
apply fast |
|
178 |
done |
|
179 |
||
58254 | 180 |
lemma BodyN: "(insert ({P}. BODY pn .{Q}) G) |-{P}. the (body pn) .{Q} ==> |
19803 | 181 |
G|-{P}. BODY pn .{Q}" |
182 |
apply (rule Body1) |
|
183 |
apply (rule_tac [2] singletonI) |
|
184 |
apply clarsimp |
|
185 |
done |
|
186 |
||
187 |
lemma escape: "[| !Z s. P Z s --> G|-{%Z s'. s'=s}.c.{%Z'. Q Z} |] ==> G|-{P}.c.{Q}" |
|
188 |
apply (rule hoare_derivs.conseq) |
|
189 |
apply fast |
|
190 |
done |
|
191 |
||
52137
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
haftmann
parents:
51717
diff
changeset
|
192 |
lemma "constant": "[| C ==> G|-{P}.c.{Q} |] ==> G|-{%Z s. P Z s & C}.c.{Q}" |
19803 | 193 |
apply (rule hoare_derivs.conseq) |
194 |
apply fast |
|
195 |
done |
|
196 |
||
197 |
lemma LoopF: "G|-{%Z s. P Z s & ~b s}.WHILE b DO c.{P}" |
|
198 |
apply (rule hoare_derivs.Loop [THEN conseq2]) |
|
199 |
apply (simp_all (no_asm)) |
|
200 |
apply (rule hoare_derivs.conseq) |
|
201 |
apply fast |
|
202 |
done |
|
203 |
||
204 |
(* |
|
205 |
Goal "[| G'||-ts; G' <= G |] ==> G||-ts" |
|
206 |
by (etac hoare_derivs.cut 1); |
|
207 |
by (etac hoare_derivs.asm 1); |
|
208 |
qed "thin"; |
|
209 |
*) |
|
210 |
lemma thin [rule_format]: "G'||-ts ==> !G. G' <= G --> G||-ts" |
|
211 |
apply (erule hoare_derivs.induct) |
|
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
212 |
apply (tactic {* ALLGOALS (EVERY'[clarify_tac @{context}, REPEAT o smp_tac @{context} 1]) *}) |
19803 | 213 |
apply (rule hoare_derivs.empty) |
214 |
apply (erule (1) hoare_derivs.insert) |
|
215 |
apply (fast intro: hoare_derivs.asm) |
|
216 |
apply (fast intro: hoare_derivs.cut) |
|
217 |
apply (fast intro: hoare_derivs.weaken) |
|
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
218 |
apply (rule hoare_derivs.conseq, intro strip, tactic "smp_tac @{context} 2 1", clarify, tactic "smp_tac @{context} 1 1",rule exI, rule exI, erule (1) conjI) |
19803 | 219 |
prefer 7 |
220 |
apply (rule_tac hoare_derivs.Body, drule_tac spec, erule_tac mp, fast) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
221 |
apply (tactic {* ALLGOALS (resolve_tac @{context} ((funpow 5 tl) @{thms hoare_derivs.intros}) THEN_ALL_NEW (fast_tac @{context})) *}) |
19803 | 222 |
done |
223 |
||
224 |
lemma weak_Body: "G|-{P}. the (body pn) .{Q} ==> G|-{P}. BODY pn .{Q}" |
|
225 |
apply (rule BodyN) |
|
226 |
apply (erule thin) |
|
227 |
apply auto |
|
228 |
done |
|
229 |
||
230 |
lemma derivs_insertD: "G||-insert t ts ==> G|-t & G||-ts" |
|
231 |
apply (fast intro: hoare_derivs.weaken) |
|
232 |
done |
|
233 |
||
58254 | 234 |
lemma finite_pointwise [rule_format (no_asm)]: "[| finite U; |
235 |
!p. G |- {P' p}.c0 p.{Q' p} --> G |- {P p}.c0 p.{Q p} |] ==> |
|
19803 | 236 |
G||-(%p. {P' p}.c0 p.{Q' p}) ` U --> G||-(%p. {P p}.c0 p.{Q p}) ` U" |
237 |
apply (erule finite_induct) |
|
238 |
apply simp |
|
239 |
apply clarsimp |
|
240 |
apply (drule derivs_insertD) |
|
241 |
apply (rule hoare_derivs.insert) |
|
242 |
apply auto |
|
243 |
done |
|
244 |
||
245 |
||
246 |
subsection "soundness" |
|
247 |
||
58254 | 248 |
lemma Loop_sound_lemma: |
249 |
"G|={P &> b}. c .{P} ==> |
|
19803 | 250 |
G|={P}. WHILE b DO c .{P &> (Not o b)}" |
251 |
apply (unfold hoare_valids_def) |
|
252 |
apply (simp (no_asm_use) add: triple_valid_def2) |
|
253 |
apply (rule allI) |
|
254 |
apply (subgoal_tac "!d s s'. <d,s> -n-> s' --> d = WHILE b DO c --> ||=n:G --> (!Z. P Z s --> P Z s' & ~b s') ") |
|
255 |
apply (erule thin_rl, fast) |
|
256 |
apply ((rule allI)+, rule impI) |
|
257 |
apply (erule evaln.induct) |
|
258 |
apply (simp_all (no_asm)) |
|
259 |
apply fast |
|
260 |
apply fast |
|
261 |
done |
|
262 |
||
58254 | 263 |
lemma Body_sound_lemma: |
264 |
"[| G Un (%pn. {P pn}. BODY pn .{Q pn})`Procs |
|
265 |
||=(%pn. {P pn}. the (body pn) .{Q pn})`Procs |] ==> |
|
19803 | 266 |
G||=(%pn. {P pn}. BODY pn .{Q pn})`Procs" |
267 |
apply (unfold hoare_valids_def) |
|
268 |
apply (rule allI) |
|
269 |
apply (induct_tac n) |
|
270 |
apply (fast intro: Body_triple_valid_0) |
|
271 |
apply clarsimp |
|
272 |
apply (drule triples_valid_Suc) |
|
273 |
apply (erule (1) notE impE) |
|
274 |
apply (simp add: ball_Un) |
|
275 |
apply (drule spec, erule impE, erule conjI, assumption) |
|
276 |
apply (fast intro!: Body_triple_valid_Suc [THEN iffD1]) |
|
277 |
done |
|
278 |
||
279 |
lemma hoare_sound: "G||-ts ==> G||=ts" |
|
280 |
apply (erule hoare_derivs.induct) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
281 |
apply (tactic {* TRYALL (eresolve_tac @{context} [@{thm Loop_sound_lemma}, @{thm Body_sound_lemma}] THEN_ALL_NEW atac) *}) |
19803 | 282 |
apply (unfold hoare_valids_def) |
283 |
apply blast |
|
284 |
apply blast |
|
285 |
apply (blast) (* asm *) |
|
286 |
apply (blast) (* cut *) |
|
287 |
apply (blast) (* weaken *) |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
23894
diff
changeset
|
288 |
apply (tactic {* ALLGOALS (EVERY' |
27239 | 289 |
[REPEAT o thin_tac @{context} "hoare_derivs ?x ?y", |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
290 |
simp_tac @{context}, clarify_tac @{context}, REPEAT o smp_tac @{context} 1]) *}) |
19803 | 291 |
apply (simp_all (no_asm_use) add: triple_valid_def2) |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
292 |
apply (intro strip, tactic "smp_tac @{context} 2 1", blast) (* conseq *) |
42793 | 293 |
apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *}) (* Skip, Ass, Local *) |
19803 | 294 |
prefer 3 apply (force) (* Call *) |
295 |
apply (erule_tac [2] evaln_elim_cases) (* If *) |
|
296 |
apply blast+ |
|
297 |
done |
|
298 |
||
299 |
||
300 |
section "completeness" |
|
301 |
||
302 |
(* Both versions *) |
|
303 |
||
304 |
(*unused*) |
|
58254 | 305 |
lemma MGT_alternI: "G|-MGT c ==> |
19803 | 306 |
G|-{%Z s0. !s1. <c,s0> -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1}" |
307 |
apply (unfold MGT_def) |
|
308 |
apply (erule conseq12) |
|
309 |
apply auto |
|
310 |
done |
|
311 |
||
312 |
(* requires com_det *) |
|
58254 | 313 |
lemma MGT_alternD: "state_not_singleton ==> |
19803 | 314 |
G|-{%Z s0. !s1. <c,s0> -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1} ==> G|-MGT c" |
315 |
apply (unfold MGT_def) |
|
316 |
apply (erule conseq12) |
|
317 |
apply auto |
|
318 |
apply (case_tac "? t. <c,?s> -c-> t") |
|
319 |
apply (fast elim: com_det) |
|
320 |
apply clarsimp |
|
321 |
apply (drule single_stateE) |
|
322 |
apply blast |
|
323 |
done |
|
324 |
||
58254 | 325 |
lemma MGF_complete: |
19803 | 326 |
"{}|-(MGT c::state triple) ==> {}|={P}.c.{Q} ==> {}|-{P}.c.{Q::state assn}" |
327 |
apply (unfold MGT_def) |
|
328 |
apply (erule conseq12) |
|
329 |
apply (clarsimp simp add: hoare_valids_def eval_eq triple_valid_def2) |
|
330 |
done |
|
331 |
||
332 |
declare WTs_elim_cases [elim!] |
|
333 |
declare not_None_eq [iff] |
|
334 |
(* requires com_det, escape (i.e. hoare_derivs.conseq) *) |
|
58254 | 335 |
lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton ==> |
19803 | 336 |
!pn:dom body. G|-{=}.BODY pn.{->} ==> WT c --> G|-{=}.c.{->}" |
337 |
apply (induct_tac c) |
|
42793 | 338 |
apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *}) |
19803 | 339 |
prefer 7 apply (fast intro: domI) |
340 |
apply (erule_tac [6] MGT_alternD) |
|
341 |
apply (unfold MGT_def) |
|
342 |
apply (drule_tac [7] bspec, erule_tac [7] domI) |
|
42793 | 343 |
apply (rule_tac [7] escape, tactic {* clarsimp_tac @{context} 7 *}, |
58254 | 344 |
rename_tac [7] "fun" y Z, |
19803 | 345 |
rule_tac [7] P1 = "%Z' s. s= (setlocs Z newlocs) [Loc Arg ::= fun Z]" in hoare_derivs.Call [THEN conseq1], erule_tac [7] conseq12) |
346 |
apply (erule_tac [!] thin_rl) |
|
347 |
apply (rule hoare_derivs.Skip [THEN conseq2]) |
|
348 |
apply (rule_tac [2] hoare_derivs.Ass [THEN conseq1]) |
|
42793 | 349 |
apply (rule_tac [3] escape, tactic {* clarsimp_tac @{context} 3 *}, |
58254 | 350 |
rename_tac [3] loc "fun" y Z, |
19803 | 351 |
rule_tac [3] P1 = "%Z' s. s= (Z[Loc loc::=fun Z])" in hoare_derivs.Local [THEN conseq1], |
352 |
erule_tac [3] conseq12) |
|
353 |
apply (erule_tac [5] hoare_derivs.Comp, erule_tac [5] conseq12) |
|
39159 | 354 |
apply (tactic {* (rtac @{thm hoare_derivs.If} THEN_ALL_NEW etac @{thm conseq12}) 6 *}) |
19803 | 355 |
apply (rule_tac [8] hoare_derivs.Loop [THEN conseq2], erule_tac [8] conseq12) |
356 |
apply auto |
|
357 |
done |
|
358 |
||
359 |
(* Version: nested single recursion *) |
|
360 |
||
361 |
lemma nesting_lemma [rule_format]: |
|
362 |
assumes "!!G ts. ts <= G ==> P G ts" |
|
363 |
and "!!G pn. P (insert (mgt_call pn) G) {mgt(the(body pn))} ==> P G {mgt_call pn}" |
|
364 |
and "!!G c. [| wt c; !pn:U. P G {mgt_call pn} |] ==> P G {mgt c}" |
|
365 |
and "!!pn. pn : U ==> wt (the (body pn))" |
|
58254 | 366 |
shows "finite U ==> uG = mgt_call`U ==> |
19803 | 367 |
!G. G <= uG --> n <= card uG --> card G = card uG - n --> (!c. wt c --> P G {mgt c})" |
368 |
apply (induct_tac n) |
|
42793 | 369 |
apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *}) |
19803 | 370 |
apply (subgoal_tac "G = mgt_call ` U") |
371 |
prefer 2 |
|
40786
0a54cfc9add3
gave more standard finite set rules simp and intro attribute
nipkow
parents:
39159
diff
changeset
|
372 |
apply (simp add: card_seteq) |
19803 | 373 |
apply simp |
41529 | 374 |
apply (erule assms(3-)) (*MGF_lemma1*) |
19803 | 375 |
apply (rule ballI) |
41529 | 376 |
apply (rule assms) (*hoare_derivs.asm*) |
19803 | 377 |
apply fast |
41529 | 378 |
apply (erule assms(3-)) (*MGF_lemma1*) |
19803 | 379 |
apply (rule ballI) |
380 |
apply (case_tac "mgt_call pn : G") |
|
41529 | 381 |
apply (rule assms) (*hoare_derivs.asm*) |
19803 | 382 |
apply fast |
41529 | 383 |
apply (rule assms(2-)) (*MGT_BodyN*) |
19803 | 384 |
apply (drule spec, erule impE, erule_tac [2] impE, drule_tac [3] spec, erule_tac [3] mp) |
41529 | 385 |
apply (erule_tac [3] assms(4-)) |
19803 | 386 |
apply fast |
387 |
apply (drule finite_subset) |
|
388 |
apply (erule finite_imageI) |
|
389 |
apply (simp (no_asm_simp)) |
|
390 |
done |
|
391 |
||
58254 | 392 |
lemma MGT_BodyN: "insert ({=}.BODY pn.{->}) G|-{=}. the (body pn) .{->} ==> |
19803 | 393 |
G|-{=}.BODY pn.{->}" |
394 |
apply (unfold MGT_def) |
|
395 |
apply (rule BodyN) |
|
396 |
apply (erule conseq2) |
|
397 |
apply force |
|
398 |
done |
|
399 |
||
400 |
(* requires BodyN, com_det *) |
|
401 |
lemma MGF: "[| state_not_singleton; WT_bodies; WT c |] ==> {}|-MGT c" |
|
402 |
apply (rule_tac P = "%G ts. G||-ts" and U = "dom body" in nesting_lemma) |
|
403 |
apply (erule hoare_derivs.asm) |
|
404 |
apply (erule MGT_BodyN) |
|
405 |
apply (rule_tac [3] finite_dom_body) |
|
406 |
apply (erule MGF_lemma1) |
|
407 |
prefer 2 apply (assumption) |
|
408 |
apply blast |
|
409 |
apply clarsimp |
|
410 |
apply (erule (1) WT_bodiesD) |
|
411 |
apply (rule_tac [3] le_refl) |
|
412 |
apply auto |
|
413 |
done |
|
414 |
||
415 |
||
416 |
(* Version: simultaneous recursion in call rule *) |
|
417 |
||
418 |
(* finiteness not really necessary here *) |
|
58254 | 419 |
lemma MGT_Body: "[| G Un (%pn. {=}. BODY pn .{->})`Procs |
420 |
||-(%pn. {=}. the (body pn) .{->})`Procs; |
|
19803 | 421 |
finite Procs |] ==> G ||-(%pn. {=}. BODY pn .{->})`Procs" |
422 |
apply (unfold MGT_def) |
|
423 |
apply (rule hoare_derivs.Body) |
|
424 |
apply (erule finite_pointwise) |
|
425 |
prefer 2 apply (assumption) |
|
426 |
apply clarify |
|
427 |
apply (erule conseq2) |
|
428 |
apply auto |
|
429 |
done |
|
430 |
||
431 |
(* requires empty, insert, com_det *) |
|
58254 | 432 |
lemma MGF_lemma2_simult [rule_format (no_asm)]: "[| state_not_singleton; WT_bodies; |
433 |
F<=(%pn. {=}.the (body pn).{->})`dom body |] ==> |
|
19803 | 434 |
(%pn. {=}. BODY pn .{->})`dom body||-F" |
435 |
apply (frule finite_subset) |
|
436 |
apply (rule finite_dom_body [THEN finite_imageI]) |
|
437 |
apply (rotate_tac 2) |
|
438 |
apply (tactic "make_imp_tac 1") |
|
439 |
apply (erule finite_induct) |
|
440 |
apply (clarsimp intro!: hoare_derivs.empty) |
|
441 |
apply (clarsimp intro!: hoare_derivs.insert simp del: range_composition) |
|
442 |
apply (erule MGF_lemma1) |
|
443 |
prefer 2 apply (fast dest: WT_bodiesD) |
|
444 |
apply clarsimp |
|
445 |
apply (rule hoare_derivs.asm) |
|
446 |
apply (fast intro: domI) |
|
447 |
done |
|
448 |
||
449 |
(* requires Body, empty, insert, com_det *) |
|
450 |
lemma MGF': "[| state_not_singleton; WT_bodies; WT c |] ==> {}|-MGT c" |
|
451 |
apply (rule MGF_lemma1) |
|
452 |
apply assumption |
|
453 |
prefer 2 apply (assumption) |
|
454 |
apply clarsimp |
|
455 |
apply (subgoal_tac "{}||- (%pn. {=}. BODY pn .{->}) `dom body") |
|
456 |
apply (erule hoare_derivs.weaken) |
|
457 |
apply (fast intro: domI) |
|
458 |
apply (rule finite_dom_body [THEN [2] MGT_Body]) |
|
459 |
apply (simp (no_asm)) |
|
460 |
apply (erule (1) MGF_lemma2_simult) |
|
461 |
apply (rule subset_refl) |
|
462 |
done |
|
463 |
||
464 |
(* requires Body+empty+insert / BodyN, com_det *) |
|
45605 | 465 |
lemmas hoare_complete = MGF' [THEN MGF_complete] |
19803 | 466 |
|
467 |
||
468 |
subsection "unused derived rules" |
|
469 |
||
470 |
lemma falseE: "G|-{%Z s. False}.c.{Q}" |
|
471 |
apply (rule hoare_derivs.conseq) |
|
472 |
apply fast |
|
473 |
done |
|
474 |
||
475 |
lemma trueI: "G|-{P}.c.{%Z s. True}" |
|
476 |
apply (rule hoare_derivs.conseq) |
|
477 |
apply (fast intro!: falseE) |
|
478 |
done |
|
479 |
||
58254 | 480 |
lemma disj: "[| G|-{P}.c.{Q}; G|-{P'}.c.{Q'} |] |
19803 | 481 |
==> G|-{%Z s. P Z s | P' Z s}.c.{%Z s. Q Z s | Q' Z s}" |
482 |
apply (rule hoare_derivs.conseq) |
|
483 |
apply (fast elim: conseq12) |
|
484 |
done (* analogue conj non-derivable *) |
|
485 |
||
486 |
lemma hoare_SkipI: "(!Z s. P Z s --> Q Z s) ==> G|-{P}. SKIP .{Q}" |
|
487 |
apply (rule conseq12) |
|
488 |
apply (rule hoare_derivs.Skip) |
|
489 |
apply fast |
|
490 |
done |
|
491 |
||
492 |
||
493 |
subsection "useful derived rules" |
|
494 |
||
495 |
lemma single_asm: "{t}|-t" |
|
496 |
apply (rule hoare_derivs.asm) |
|
497 |
apply (rule subset_refl) |
|
498 |
done |
|
499 |
||
500 |
lemma export_s: "[| !!s'. G|-{%Z s. s'=s & P Z s}.c.{Q} |] ==> G|-{P}.c.{Q}" |
|
501 |
apply (rule hoare_derivs.conseq) |
|
502 |
apply auto |
|
503 |
done |
|
504 |
||
505 |
||
58254 | 506 |
lemma weak_Local: "[| G|-{P}. c .{Q}; !k Z s. Q Z s --> Q Z (s[Loc Y::=k]) |] ==> |
19803 | 507 |
G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{Q}" |
508 |
apply (rule export_s) |
|
509 |
apply (rule hoare_derivs.Local) |
|
510 |
apply (erule conseq2) |
|
511 |
apply (erule spec) |
|
512 |
done |
|
513 |
||
514 |
(* |
|
515 |
Goal "!Q. G |-{%Z s. ~(? s'. <c,s> -c-> s')}. c .{Q}" |
|
516 |
by (induct_tac "c" 1); |
|
517 |
by Auto_tac; |
|
518 |
by (rtac conseq1 1); |
|
519 |
by (rtac hoare_derivs.Skip 1); |
|
520 |
force 1; |
|
521 |
by (rtac conseq1 1); |
|
522 |
by (rtac hoare_derivs.Ass 1); |
|
523 |
force 1; |
|
524 |
by (defer_tac 1); |
|
525 |
### |
|
526 |
by (rtac hoare_derivs.Comp 1); |
|
527 |
by (dtac spec 2); |
|
528 |
by (dtac spec 2); |
|
529 |
by (assume_tac 2); |
|
530 |
by (etac conseq1 2); |
|
531 |
by (Clarsimp_tac 2); |
|
532 |
force 1; |
|
533 |
*) |
|
8177 | 534 |
|
535 |
end |