author | blanchet |
Mon, 31 Jan 2022 16:09:23 +0100 | |
changeset 75020 | b087610592b4 |
parent 72985 | 9cc431444435 |
child 76987 | 4c275405faae |
permissions | -rw-r--r-- |
33026 | 1 |
(* Title: HOL/Isar_Examples/Hoare.thy |
61932 | 2 |
Author: Makarius |
10148 | 3 |
|
4 |
A formulation of Hoare logic suitable for Isar. |
|
5 |
*) |
|
6 |
||
58882 | 7 |
section \<open>Hoare Logic\<close> |
10148 | 8 |
|
31758 | 9 |
theory Hoare |
72985
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
10 |
imports "HOL-Hoare.Hoare_Tac" |
31758 | 11 |
begin |
10148 | 12 |
|
58614 | 13 |
subsection \<open>Abstract syntax and semantics\<close> |
10148 | 14 |
|
61932 | 15 |
text \<open> |
16 |
The following abstract syntax and semantics of Hoare Logic over \<^verbatim>\<open>WHILE\<close> |
|
17 |
programs closely follows the existing tradition in Isabelle/HOL of |
|
18 |
formalizing the presentation given in @{cite \<open>\S6\<close> "Winskel:1993"}. See also |
|
63680 | 19 |
\<^dir>\<open>~~/src/HOL/Hoare\<close> and @{cite "Nipkow:1998:Winskel"}. |
61932 | 20 |
\<close> |
10148 | 21 |
|
41818 | 22 |
type_synonym 'a bexp = "'a set" |
23 |
type_synonym 'a assn = "'a set" |
|
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
24 |
type_synonym 'a var = "'a \<Rightarrow> nat" |
10148 | 25 |
|
58310 | 26 |
datatype 'a com = |
55656 | 27 |
Basic "'a \<Rightarrow> 'a" |
10148 | 28 |
| Seq "'a com" "'a com" ("(_;/ _)" [60, 61] 60) |
29 |
| Cond "'a bexp" "'a com" "'a com" |
|
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
30 |
| While "'a bexp" "'a assn" "'a var" "'a com" |
10148 | 31 |
|
37671 | 32 |
abbreviation Skip ("SKIP") |
55656 | 33 |
where "SKIP \<equiv> Basic id" |
10148 | 34 |
|
55656 | 35 |
type_synonym 'a sem = "'a \<Rightarrow> 'a \<Rightarrow> bool" |
10148 | 36 |
|
55656 | 37 |
primrec iter :: "nat \<Rightarrow> 'a bexp \<Rightarrow> 'a sem \<Rightarrow> 'a sem" |
63585 | 38 |
where |
39 |
"iter 0 b S s s' \<longleftrightarrow> s \<notin> b \<and> s = s'" |
|
40 |
| "iter (Suc n) b S s s' \<longleftrightarrow> s \<in> b \<and> (\<exists>s''. S s s'' \<and> iter n b S s'' s')" |
|
10148 | 41 |
|
55656 | 42 |
primrec Sem :: "'a com \<Rightarrow> 'a sem" |
63585 | 43 |
where |
44 |
"Sem (Basic f) s s' \<longleftrightarrow> s' = f s" |
|
45 |
| "Sem (c1; c2) s s' \<longleftrightarrow> (\<exists>s''. Sem c1 s s'' \<and> Sem c2 s'' s')" |
|
46 |
| "Sem (Cond b c1 c2) s s' \<longleftrightarrow> (if s \<in> b then Sem c1 s s' else Sem c2 s s')" |
|
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
47 |
| "Sem (While b x y c) s s' \<longleftrightarrow> (\<exists>n. iter n b (Sem c) s s')" |
10148 | 48 |
|
63585 | 49 |
definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50) |
55656 | 50 |
where "\<turnstile> P c Q \<longleftrightarrow> (\<forall>s s'. Sem c s s' \<longrightarrow> s \<in> P \<longrightarrow> s' \<in> Q)" |
10148 | 51 |
|
63585 | 52 |
lemma ValidI [intro?]: "(\<And>s s'. Sem c s s' \<Longrightarrow> s \<in> P \<Longrightarrow> s' \<in> Q) \<Longrightarrow> \<turnstile> P c Q" |
10148 | 53 |
by (simp add: Valid_def) |
54 |
||
63585 | 55 |
lemma ValidD [dest?]: "\<turnstile> P c Q \<Longrightarrow> Sem c s s' \<Longrightarrow> s \<in> P \<Longrightarrow> s' \<in> Q" |
10148 | 56 |
by (simp add: Valid_def) |
57 |
||
58 |
||
58614 | 59 |
subsection \<open>Primitive Hoare rules\<close> |
10148 | 60 |
|
61932 | 61 |
text \<open> |
62 |
From the semantics defined above, we derive the standard set of primitive |
|
63 |
Hoare rules; e.g.\ see @{cite \<open>\S6\<close> "Winskel:1993"}. Usually, variant forms |
|
64 |
of these rules are applied in actual proof, see also \S\ref{sec:hoare-isar} |
|
65 |
and \S\ref{sec:hoare-vcg}. |
|
10148 | 66 |
|
61541 | 67 |
\<^medskip> |
68 |
The \<open>basic\<close> rule represents any kind of atomic access to the state space. |
|
69 |
This subsumes the common rules of \<open>skip\<close> and \<open>assign\<close>, as formulated in |
|
61932 | 70 |
\S\ref{sec:hoare-isar}. |
71 |
\<close> |
|
10148 | 72 |
|
55656 | 73 |
theorem basic: "\<turnstile> {s. f s \<in> P} (Basic f) P" |
10148 | 74 |
proof |
55656 | 75 |
fix s s' |
76 |
assume s: "s \<in> {s. f s \<in> P}" |
|
10148 | 77 |
assume "Sem (Basic f) s s'" |
37671 | 78 |
then have "s' = f s" by simp |
55656 | 79 |
with s show "s' \<in> P" by simp |
10148 | 80 |
qed |
81 |
||
61932 | 82 |
text \<open> |
83 |
The rules for sequential commands and semantic consequences are established |
|
84 |
in a straight forward manner as follows. |
|
85 |
\<close> |
|
10148 | 86 |
|
55656 | 87 |
theorem seq: "\<turnstile> P c1 Q \<Longrightarrow> \<turnstile> Q c2 R \<Longrightarrow> \<turnstile> P (c1; c2) R" |
10148 | 88 |
proof |
55656 | 89 |
assume cmd1: "\<turnstile> P c1 Q" and cmd2: "\<turnstile> Q c2 R" |
90 |
fix s s' |
|
91 |
assume s: "s \<in> P" |
|
10148 | 92 |
assume "Sem (c1; c2) s s'" |
93 |
then obtain s'' where sem1: "Sem c1 s s''" and sem2: "Sem c2 s'' s'" |
|
94 |
by auto |
|
55656 | 95 |
from cmd1 sem1 s have "s'' \<in> Q" .. |
96 |
with cmd2 sem2 show "s' \<in> R" .. |
|
10148 | 97 |
qed |
98 |
||
55656 | 99 |
theorem conseq: "P' \<subseteq> P \<Longrightarrow> \<turnstile> P c Q \<Longrightarrow> Q \<subseteq> Q' \<Longrightarrow> \<turnstile> P' c Q'" |
10148 | 100 |
proof |
55656 | 101 |
assume P'P: "P' \<subseteq> P" and QQ': "Q \<subseteq> Q'" |
102 |
assume cmd: "\<turnstile> P c Q" |
|
10148 | 103 |
fix s s' :: 'a |
104 |
assume sem: "Sem c s s'" |
|
67613 | 105 |
assume "s \<in> P'" with P'P have "s \<in> P" .. |
55656 | 106 |
with cmd sem have "s' \<in> Q" .. |
107 |
with QQ' show "s' \<in> Q'" .. |
|
10148 | 108 |
qed |
109 |
||
61932 | 110 |
text \<open> |
111 |
The rule for conditional commands is directly reflected by the corresponding |
|
112 |
semantics; in the proof we just have to look closely which cases apply. |
|
113 |
\<close> |
|
10148 | 114 |
|
115 |
theorem cond: |
|
55656 | 116 |
assumes case_b: "\<turnstile> (P \<inter> b) c1 Q" |
117 |
and case_nb: "\<turnstile> (P \<inter> -b) c2 Q" |
|
118 |
shows "\<turnstile> P (Cond b c1 c2) Q" |
|
10148 | 119 |
proof |
55656 | 120 |
fix s s' |
121 |
assume s: "s \<in> P" |
|
10148 | 122 |
assume sem: "Sem (Cond b c1 c2) s s'" |
55656 | 123 |
show "s' \<in> Q" |
10148 | 124 |
proof cases |
55656 | 125 |
assume b: "s \<in> b" |
10148 | 126 |
from case_b show ?thesis |
127 |
proof |
|
128 |
from sem b show "Sem c1 s s'" by simp |
|
55656 | 129 |
from s b show "s \<in> P \<inter> b" by simp |
10148 | 130 |
qed |
131 |
next |
|
55656 | 132 |
assume nb: "s \<notin> b" |
10148 | 133 |
from case_nb show ?thesis |
134 |
proof |
|
135 |
from sem nb show "Sem c2 s s'" by simp |
|
67613 | 136 |
from s nb show "s \<in> P \<inter> -b" by simp |
10148 | 137 |
qed |
138 |
qed |
|
139 |
qed |
|
140 |
||
61932 | 141 |
text \<open> |
142 |
The \<open>while\<close> rule is slightly less trivial --- it is the only one based on |
|
143 |
recursion, which is expressed in the semantics by a Kleene-style least |
|
61541 | 144 |
fixed-point construction. The auxiliary statement below, which is by |
145 |
induction on the number of iterations is the main point to be proven; the |
|
61932 | 146 |
rest is by routine application of the semantics of \<^verbatim>\<open>WHILE\<close>. |
147 |
\<close> |
|
10148 | 148 |
|
18241 | 149 |
theorem while: |
55656 | 150 |
assumes body: "\<turnstile> (P \<inter> b) c P" |
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
151 |
shows "\<turnstile> P (While b X Y c) (P \<inter> -b)" |
10148 | 152 |
proof |
55656 | 153 |
fix s s' assume s: "s \<in> P" |
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
154 |
assume "Sem (While b X Y c) s s'" |
18241 | 155 |
then obtain n where "iter n b (Sem c) s s'" by auto |
55656 | 156 |
from this and s show "s' \<in> P \<inter> -b" |
20503 | 157 |
proof (induct n arbitrary: s) |
19122 | 158 |
case 0 |
37671 | 159 |
then show ?case by auto |
11987 | 160 |
next |
19122 | 161 |
case (Suc n) |
55656 | 162 |
then obtain s'' where b: "s \<in> b" and sem: "Sem c s s''" |
37671 | 163 |
and iter: "iter n b (Sem c) s'' s'" by auto |
55656 | 164 |
from Suc and b have "s \<in> P \<inter> b" by simp |
165 |
with body sem have "s'' \<in> P" .. |
|
11987 | 166 |
with iter show ?case by (rule Suc) |
10148 | 167 |
qed |
168 |
qed |
|
169 |
||
170 |
||
58614 | 171 |
subsection \<open>Concrete syntax for assertions\<close> |
10148 | 172 |
|
61932 | 173 |
text \<open> |
174 |
We now introduce concrete syntax for describing commands (with embedded |
|
175 |
expressions) and assertions. The basic technique is that of semantic |
|
176 |
``quote-antiquote''. A \<^emph>\<open>quotation\<close> is a syntactic entity delimited by an |
|
177 |
implicit abstraction, say over the state space. An \<^emph>\<open>antiquotation\<close> is a |
|
178 |
marked expression within a quotation that refers the implicit argument; a |
|
179 |
typical antiquotation would select (or even update) components from the |
|
180 |
state. |
|
10148 | 181 |
|
61541 | 182 |
We will see some examples later in the concrete rules and applications. |
10148 | 183 |
|
61541 | 184 |
\<^medskip> |
185 |
The following specification of syntax and translations is for Isabelle |
|
186 |
experts only; feel free to ignore it. |
|
10148 | 187 |
|
61541 | 188 |
While the first part is still a somewhat intelligible specification of the |
189 |
concrete syntactic representation of our Hoare language, the actual ``ML |
|
190 |
drivers'' is quite involved. Just note that the we re-use the basic |
|
69597 | 191 |
quote/antiquote translations as already defined in Isabelle/Pure (see \<^ML>\<open>Syntax_Trans.quote_tr\<close>, and \<^ML>\<open>Syntax_Trans.quote_tr'\<close>,). |
61932 | 192 |
\<close> |
10148 | 193 |
|
194 |
syntax |
|
55662 | 195 |
"_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)" |
196 |
"_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" ("\<acute>_" [1000] 1000) |
|
197 |
"_Subst" :: "'a bexp \<Rightarrow> 'b \<Rightarrow> idt \<Rightarrow> 'a bexp" ("_[_'/\<acute>_]" [1000] 999) |
|
198 |
"_Assert" :: "'a \<Rightarrow> 'a set" ("(\<lbrace>_\<rbrace>)" [0] 1000) |
|
199 |
"_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" ("(\<acute>_ :=/ _)" [70, 65] 61) |
|
200 |
"_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" |
|
201 |
("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) |
|
202 |
"_While_inv" :: "'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a com \<Rightarrow> 'a com" |
|
203 |
("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) |
|
204 |
"_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0WHILE _ //DO _ /OD)" [0, 0] 61) |
|
10148 | 205 |
|
206 |
translations |
|
55662 | 207 |
"\<lbrace>b\<rbrace>" \<rightharpoonup> "CONST Collect (_quote b)" |
208 |
"B [a/\<acute>x]" \<rightharpoonup> "\<lbrace>\<acute>(_update_name x (\<lambda>_. a)) \<in> B\<rbrace>" |
|
209 |
"\<acute>x := a" \<rightharpoonup> "CONST Basic (_quote (\<acute>(_update_name x (\<lambda>_. a))))" |
|
55656 | 210 |
"IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "CONST Cond \<lbrace>b\<rbrace> c1 c2" |
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
211 |
"WHILE b INV i DO c OD" \<rightharpoonup> "CONST While \<lbrace>b\<rbrace> i (\<lambda>_. 0) c" |
55662 | 212 |
"WHILE b DO c OD" \<rightleftharpoons> "WHILE b INV CONST undefined DO c OD" |
10148 | 213 |
|
58614 | 214 |
parse_translation \<open> |
10148 | 215 |
let |
69597 | 216 |
fun quote_tr [t] = Syntax_Trans.quote_tr \<^syntax_const>\<open>_antiquote\<close> t |
10148 | 217 |
| quote_tr ts = raise TERM ("quote_tr", ts); |
69597 | 218 |
in [(\<^syntax_const>\<open>_quote\<close>, K quote_tr)] end |
58614 | 219 |
\<close> |
10148 | 220 |
|
61932 | 221 |
text \<open> |
222 |
As usual in Isabelle syntax translations, the part for printing is more |
|
223 |
complicated --- we cannot express parts as macro rules as above. Don't look |
|
224 |
here, unless you have to do similar things for yourself. |
|
225 |
\<close> |
|
10148 | 226 |
|
58614 | 227 |
print_translation \<open> |
10148 | 228 |
let |
229 |
fun quote_tr' f (t :: ts) = |
|
69597 | 230 |
Term.list_comb (f $ Syntax_Trans.quote_tr' \<^syntax_const>\<open>_antiquote\<close> t, ts) |
10148 | 231 |
| quote_tr' _ _ = raise Match; |
232 |
||
69597 | 233 |
val assert_tr' = quote_tr' (Syntax.const \<^syntax_const>\<open>_Assert\<close>); |
10148 | 234 |
|
69597 | 235 |
fun bexp_tr' name ((Const (\<^const_syntax>\<open>Collect\<close>, _) $ t) :: ts) = |
10148 | 236 |
quote_tr' (Syntax.const name) (t :: ts) |
237 |
| bexp_tr' _ _ = raise Match; |
|
238 |
||
25706 | 239 |
fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = |
69597 | 240 |
quote_tr' (Syntax.const \<^syntax_const>\<open>_Assign\<close> $ Syntax_Trans.update_name_tr' f) |
42284 | 241 |
(Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts) |
10148 | 242 |
| assign_tr' _ = raise Match; |
243 |
in |
|
69597 | 244 |
[(\<^const_syntax>\<open>Collect\<close>, K assert_tr'), |
245 |
(\<^const_syntax>\<open>Basic\<close>, K assign_tr'), |
|
246 |
(\<^const_syntax>\<open>Cond\<close>, K (bexp_tr' \<^syntax_const>\<open>_Cond\<close>)), |
|
247 |
(\<^const_syntax>\<open>While\<close>, K (bexp_tr' \<^syntax_const>\<open>_While_inv\<close>))] |
|
10148 | 248 |
end |
58614 | 249 |
\<close> |
10148 | 250 |
|
251 |
||
58614 | 252 |
subsection \<open>Rules for single-step proof \label{sec:hoare-isar}\<close> |
10148 | 253 |
|
61932 | 254 |
text \<open> |
255 |
We are now ready to introduce a set of Hoare rules to be used in single-step |
|
256 |
structured proofs in Isabelle/Isar. We refer to the concrete syntax |
|
257 |
introduce above. |
|
10148 | 258 |
|
61541 | 259 |
\<^medskip> |
260 |
Assertions of Hoare Logic may be manipulated in calculational proofs, with |
|
261 |
the inclusion expressed in terms of sets or predicates. Reversed order is |
|
61932 | 262 |
supported as well. |
263 |
\<close> |
|
10148 | 264 |
|
55656 | 265 |
lemma [trans]: "\<turnstile> P c Q \<Longrightarrow> P' \<subseteq> P \<Longrightarrow> \<turnstile> P' c Q" |
10148 | 266 |
by (unfold Valid_def) blast |
55656 | 267 |
lemma [trans] : "P' \<subseteq> P \<Longrightarrow> \<turnstile> P c Q \<Longrightarrow> \<turnstile> P' c Q" |
10148 | 268 |
by (unfold Valid_def) blast |
269 |
||
55656 | 270 |
lemma [trans]: "Q \<subseteq> Q' \<Longrightarrow> \<turnstile> P c Q \<Longrightarrow> \<turnstile> P c Q'" |
10148 | 271 |
by (unfold Valid_def) blast |
55656 | 272 |
lemma [trans]: "\<turnstile> P c Q \<Longrightarrow> Q \<subseteq> Q' \<Longrightarrow> \<turnstile> P c Q'" |
10148 | 273 |
by (unfold Valid_def) blast |
274 |
||
275 |
lemma [trans]: |
|
55656 | 276 |
"\<turnstile> \<lbrace>\<acute>P\<rbrace> c Q \<Longrightarrow> (\<And>s. P' s \<longrightarrow> P s) \<Longrightarrow> \<turnstile> \<lbrace>\<acute>P'\<rbrace> c Q" |
10148 | 277 |
by (simp add: Valid_def) |
278 |
lemma [trans]: |
|
55656 | 279 |
"(\<And>s. P' s \<longrightarrow> P s) \<Longrightarrow> \<turnstile> \<lbrace>\<acute>P\<rbrace> c Q \<Longrightarrow> \<turnstile> \<lbrace>\<acute>P'\<rbrace> c Q" |
10148 | 280 |
by (simp add: Valid_def) |
281 |
||
282 |
lemma [trans]: |
|
55656 | 283 |
"\<turnstile> P c \<lbrace>\<acute>Q\<rbrace> \<Longrightarrow> (\<And>s. Q s \<longrightarrow> Q' s) \<Longrightarrow> \<turnstile> P c \<lbrace>\<acute>Q'\<rbrace>" |
10148 | 284 |
by (simp add: Valid_def) |
285 |
lemma [trans]: |
|
55656 | 286 |
"(\<And>s. Q s \<longrightarrow> Q' s) \<Longrightarrow> \<turnstile> P c \<lbrace>\<acute>Q\<rbrace> \<Longrightarrow> \<turnstile> P c \<lbrace>\<acute>Q'\<rbrace>" |
10148 | 287 |
by (simp add: Valid_def) |
288 |
||
289 |
||
61932 | 290 |
text \<open> |
291 |
Identity and basic assignments.\<^footnote>\<open>The \<open>hoare\<close> method introduced in |
|
292 |
\S\ref{sec:hoare-vcg} is able to provide proper instances for any number of |
|
293 |
basic assignments, without producing additional verification conditions.\<close> |
|
294 |
\<close> |
|
10148 | 295 |
|
55656 | 296 |
lemma skip [intro?]: "\<turnstile> P SKIP P" |
10148 | 297 |
proof - |
55656 | 298 |
have "\<turnstile> {s. id s \<in> P} SKIP P" by (rule basic) |
37671 | 299 |
then show ?thesis by simp |
10148 | 300 |
qed |
301 |
||
55656 | 302 |
lemma assign: "\<turnstile> P [\<acute>a/\<acute>x::'a] \<acute>x := \<acute>a P" |
10148 | 303 |
by (rule basic) |
304 |
||
61932 | 305 |
text \<open> |
306 |
Note that above formulation of assignment corresponds to our preferred way |
|
307 |
to model state spaces, using (extensible) record types in HOL @{cite |
|
308 |
"Naraschewski-Wenzel:1998:HOOL"}. For any record field \<open>x\<close>, Isabelle/HOL |
|
309 |
provides a functions \<open>x\<close> (selector) and \<open>x_update\<close> (update). Above, there is |
|
310 |
only a place-holder appearing for the latter kind of function: due to |
|
311 |
concrete syntax \<open>\<acute>x := \<acute>a\<close> also contains \<open>x_update\<close>.\<^footnote>\<open>Note that due to the |
|
312 |
external nature of HOL record fields, we could not even state a general |
|
313 |
theorem relating selector and update functions (if this were required here); |
|
314 |
this would only work for any particular instance of record fields introduced |
|
315 |
so far.\<close> |
|
10148 | 316 |
|
61541 | 317 |
\<^medskip> |
61932 | 318 |
Sequential composition --- normalizing with associativity achieves proper of |
319 |
chunks of code verified separately. |
|
320 |
\<close> |
|
10148 | 321 |
|
322 |
lemmas [trans, intro?] = seq |
|
323 |
||
55656 | 324 |
lemma seq_assoc [simp]: "\<turnstile> P c1;(c2;c3) Q \<longleftrightarrow> \<turnstile> P (c1;c2);c3 Q" |
10148 | 325 |
by (auto simp add: Valid_def) |
326 |
||
58614 | 327 |
text \<open>Conditional statements.\<close> |
10148 | 328 |
|
329 |
lemmas [trans, intro?] = cond |
|
330 |
||
331 |
lemma [trans, intro?]: |
|
55656 | 332 |
"\<turnstile> \<lbrace>\<acute>P \<and> \<acute>b\<rbrace> c1 Q |
333 |
\<Longrightarrow> \<turnstile> \<lbrace>\<acute>P \<and> \<not> \<acute>b\<rbrace> c2 Q |
|
334 |
\<Longrightarrow> \<turnstile> \<lbrace>\<acute>P\<rbrace> IF \<acute>b THEN c1 ELSE c2 FI Q" |
|
10148 | 335 |
by (rule cond) (simp_all add: Valid_def) |
336 |
||
58614 | 337 |
text \<open>While statements --- with optional invariant.\<close> |
10148 | 338 |
|
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
339 |
lemma [intro?]: "\<turnstile> (P \<inter> b) c P \<Longrightarrow> \<turnstile> P (While b P V c) (P \<inter> -b)" |
10148 | 340 |
by (rule while) |
341 |
||
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
342 |
lemma [intro?]: "\<turnstile> (P \<inter> b) c P \<Longrightarrow> \<turnstile> P (While b undefined V c) (P \<inter> -b)" |
10148 | 343 |
by (rule while) |
344 |
||
345 |
||
346 |
lemma [intro?]: |
|
55656 | 347 |
"\<turnstile> \<lbrace>\<acute>P \<and> \<acute>b\<rbrace> c \<lbrace>\<acute>P\<rbrace> |
348 |
\<Longrightarrow> \<turnstile> \<lbrace>\<acute>P\<rbrace> WHILE \<acute>b INV \<lbrace>\<acute>P\<rbrace> DO c OD \<lbrace>\<acute>P \<and> \<not> \<acute>b\<rbrace>" |
|
10148 | 349 |
by (simp add: while Collect_conj_eq Collect_neg_eq) |
350 |
||
351 |
lemma [intro?]: |
|
55656 | 352 |
"\<turnstile> \<lbrace>\<acute>P \<and> \<acute>b\<rbrace> c \<lbrace>\<acute>P\<rbrace> |
353 |
\<Longrightarrow> \<turnstile> \<lbrace>\<acute>P\<rbrace> WHILE \<acute>b DO c OD \<lbrace>\<acute>P \<and> \<not> \<acute>b\<rbrace>" |
|
10148 | 354 |
by (simp add: while Collect_conj_eq Collect_neg_eq) |
355 |
||
356 |
||
58614 | 357 |
subsection \<open>Verification conditions \label{sec:hoare-vcg}\<close> |
10148 | 358 |
|
61932 | 359 |
text \<open> |
360 |
We now load the \<^emph>\<open>original\<close> ML file for proof scripts and tactic definition |
|
63680 | 361 |
for the Hoare Verification Condition Generator (see \<^dir>\<open>~~/src/HOL/Hoare\<close>). |
362 |
As far as we are concerned here, the result is a proof method \<open>hoare\<close>, which |
|
363 |
may be applied to a Hoare Logic assertion to extract purely logical |
|
364 |
verification conditions. It is important to note that the method requires |
|
365 |
\<^verbatim>\<open>WHILE\<close> loops to be fully annotated with invariants beforehand. |
|
366 |
Furthermore, only \<^emph>\<open>concrete\<close> pieces of code are handled --- the underlying |
|
367 |
tactic fails ungracefully if supplied with meta-variables or parameters, for |
|
368 |
example. |
|
61932 | 369 |
\<close> |
10148 | 370 |
|
13862 | 371 |
lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q" |
18193 | 372 |
by (auto simp add: Valid_def) |
13862 | 373 |
|
374 |
lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q" |
|
18193 | 375 |
by (auto simp: Valid_def) |
13862 | 376 |
|
377 |
lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R" |
|
18193 | 378 |
by (auto simp: Valid_def) |
13862 | 379 |
|
380 |
lemma CondRule: |
|
18193 | 381 |
"p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')} |
382 |
\<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q" |
|
383 |
by (auto simp: Valid_def) |
|
13862 | 384 |
|
18241 | 385 |
lemma iter_aux: |
55656 | 386 |
"\<forall>s s'. Sem c s s' \<longrightarrow> s \<in> I \<and> s \<in> b \<longrightarrow> s' \<in> I \<Longrightarrow> |
387 |
(\<And>s s'. s \<in> I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' \<in> I \<and> s' \<notin> b)" |
|
388 |
by (induct n) auto |
|
13862 | 389 |
|
390 |
lemma WhileRule: |
|
72806
4fa08e083865
Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents:
69605
diff
changeset
|
391 |
"p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i v c) q" |
18193 | 392 |
apply (clarsimp simp: Valid_def) |
393 |
apply (drule iter_aux) |
|
394 |
prefer 2 |
|
395 |
apply assumption |
|
396 |
apply blast |
|
397 |
apply blast |
|
398 |
done |
|
13862 | 399 |
|
72985
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
400 |
declare BasicRule [Hoare_Tac.BasicRule] |
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
401 |
and SkipRule [Hoare_Tac.SkipRule] |
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
402 |
and SeqRule [Hoare_Tac.SeqRule] |
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
403 |
and CondRule [Hoare_Tac.CondRule] |
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
404 |
and WhileRule [Hoare_Tac.WhileRule] |
10148 | 405 |
|
58614 | 406 |
method_setup hoare = |
407 |
\<open>Scan.succeed (fn ctxt => |
|
30510
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents:
28524
diff
changeset
|
408 |
(SIMPLE_METHOD' |
72985
9cc431444435
clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents:
72806
diff
changeset
|
409 |
(Hoare_Tac.hoare_tac ctxt |
58614 | 410 |
(simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm "Record.K_record_comp"}] )))))\<close> |
10148 | 411 |
"verification condition generator for Hoare logic" |
412 |
||
13703 | 413 |
end |