author | wenzelm |
Mon, 16 Aug 1999 14:22:45 +0200 | |
changeset 7208 | 8b4acb408301 |
parent 5515 | 903c956beac3 |
child 9241 | f961c1fdff50 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/IMP/Hoare.ML |
938 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Tobias Nipkow |
936 | 4 |
Copyright 1995 TUM |
5 |
||
1481 | 6 |
Soundness (and part of) relative completeness of Hoare rules |
7 |
wrt denotational semantics |
|
936 | 8 |
*) |
9 |
||
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
10 |
Goal "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"; |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
11 |
by (etac hoare.conseq 1); |
5515 | 12 |
by (atac 1); |
13 |
by (Fast_tac 1); |
|
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
14 |
qed "hoare_conseq1"; |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
15 |
|
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
16 |
Goal "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"; |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
17 |
by (rtac hoare.conseq 1); |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
18 |
by (atac 2); |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
19 |
by (ALLGOALS Fast_tac); |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
20 |
qed "hoare_conseq2"; |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
21 |
|
5117 | 22 |
Goalw [hoare_valid_def] "|- {P}c{Q} ==> |= {P}c{Q}"; |
1730 | 23 |
by (etac hoare.induct 1); |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
24 |
by (ALLGOALS Asm_simp_tac); |
1973 | 25 |
by (Fast_tac 1); |
1910 | 26 |
by (Fast_tac 1); |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
27 |
by (EVERY' [rtac allI, rtac allI, rtac impI] 1); |
1465 | 28 |
by (etac induct2 1); |
2055 | 29 |
by (rtac Gamma_mono 1); |
1465 | 30 |
by (rewtac Gamma_def); |
1973 | 31 |
by (Fast_tac 1); |
1730 | 32 |
qed "hoare_sound"; |
936 | 33 |
|
5069 | 34 |
Goalw [wp_def] "wp SKIP Q = Q"; |
2031 | 35 |
by (Simp_tac 1); |
2810 | 36 |
qed "wp_SKIP"; |
1481 | 37 |
|
5069 | 38 |
Goalw [wp_def] "wp (x:=a) Q = (%s. Q(s[x:=a s]))"; |
2031 | 39 |
by (Simp_tac 1); |
2810 | 40 |
qed "wp_Ass"; |
1481 | 41 |
|
5069 | 42 |
Goalw [wp_def] "wp (c;d) Q = wp c (wp d Q)"; |
2031 | 43 |
by (Simp_tac 1); |
44 |
by (rtac ext 1); |
|
1910 | 45 |
by (Fast_tac 1); |
2810 | 46 |
qed "wp_Semi"; |
936 | 47 |
|
5069 | 48 |
Goalw [wp_def] |
5117 | 49 |
"wp (IF b THEN c ELSE d) Q = (%s. (b s --> wp c Q s) & (~b s --> wp d Q s))"; |
2031 | 50 |
by (Simp_tac 1); |
51 |
by (rtac ext 1); |
|
1910 | 52 |
by (Fast_tac 1); |
2810 | 53 |
qed "wp_If"; |
936 | 54 |
|
5069 | 55 |
Goalw [wp_def] |
5117 | 56 |
"b s ==> wp (WHILE b DO c) Q s = wp (c;WHILE b DO c) Q s"; |
2031 | 57 |
by (stac C_While_If 1); |
58 |
by (Asm_simp_tac 1); |
|
2810 | 59 |
qed "wp_While_True"; |
1481 | 60 |
|
5117 | 61 |
Goalw [wp_def] "~b s ==> wp (WHILE b DO c) Q s = Q s"; |
2031 | 62 |
by (stac C_While_If 1); |
63 |
by (Asm_simp_tac 1); |
|
2810 | 64 |
qed "wp_While_False"; |
1481 | 65 |
|
2810 | 66 |
Addsimps [wp_SKIP,wp_Ass,wp_Semi,wp_If,wp_While_True,wp_While_False]; |
1481 | 67 |
|
1910 | 68 |
(*Not suitable for rewriting: LOOPS!*) |
5278 | 69 |
Goal "wp (WHILE b DO c) Q s = (if b s then wp (c;WHILE b DO c) Q s else Q s)"; |
4686 | 70 |
by (Simp_tac 1); |
2810 | 71 |
qed "wp_While_if"; |
1910 | 72 |
|
5278 | 73 |
Goal "wp (WHILE b DO c) Q s = \ |
3842 | 74 |
\ (s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))"; |
4686 | 75 |
by (Simp_tac 1); |
3023 | 76 |
by (rtac iffI 1); |
77 |
by (rtac weak_coinduct 1); |
|
78 |
by (etac CollectI 1); |
|
4153 | 79 |
by Safe_tac; |
3023 | 80 |
by (rotate_tac ~1 1); |
81 |
by (Asm_full_simp_tac 1); |
|
82 |
by (rotate_tac ~1 1); |
|
83 |
by (Asm_full_simp_tac 1); |
|
4089 | 84 |
by (asm_full_simp_tac (simpset() addsimps [wp_def,Gamma_def]) 1); |
3023 | 85 |
by (strip_tac 1); |
86 |
by (rtac mp 1); |
|
87 |
by (assume_tac 2); |
|
88 |
by (etac induct2 1); |
|
4089 | 89 |
by (fast_tac (claset() addSIs [monoI]) 1); |
3023 | 90 |
by (stac gfp_Tarski 1); |
4089 | 91 |
by (fast_tac (claset() addSIs [monoI]) 1); |
3023 | 92 |
by (Fast_tac 1); |
2810 | 93 |
qed "wp_While"; |
1910 | 94 |
|
1481 | 95 |
Delsimps [C_while]; |
936 | 96 |
|
1910 | 97 |
AddSIs [hoare.skip, hoare.ass, hoare.semi, hoare.If]; |
98 |
||
5069 | 99 |
Goal "!Q. |- {wp c Q} c {Q}"; |
5183 | 100 |
by (induct_tac "c" 1); |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
101 |
by (ALLGOALS Simp_tac); |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
102 |
by (REPEAT_FIRST Fast_tac); |
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
103 |
by (blast_tac (claset() addIs [hoare_conseq1]) 1); |
3737 | 104 |
by Safe_tac; |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
105 |
by (rtac hoare_conseq2 1); |
2055 | 106 |
by (rtac hoare.While 1); |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
107 |
by (rtac hoare_conseq1 1); |
1910 | 108 |
by (Fast_tac 2); |
2055 | 109 |
by (safe_tac HOL_cs); |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
110 |
by (ALLGOALS (EVERY'[rotate_tac ~1, Asm_full_simp_tac])); |
2810 | 111 |
qed_spec_mp "wp_is_pre"; |
1481 | 112 |
|
5117 | 113 |
Goal "|= {P}c{Q} ==> |- {P}c{Q}"; |
5301
e24d15594edd
streamlined proofs with new hoare_conseq1, hoare_conseq2
oheimb
parents:
5278
diff
changeset
|
114 |
by (rtac (wp_is_pre RSN (2,hoare_conseq1)) 1); |
2810 | 115 |
by (rewrite_goals_tac [hoare_valid_def,wp_def]); |
1910 | 116 |
by (Fast_tac 1); |
1481 | 117 |
qed "hoare_relative_complete"; |