| author | wenzelm |
| Wed, 04 Nov 2015 23:27:00 +0100 | |
| changeset 61578 | 6623c81cb15a |
| parent 61541 | 846c72206207 |
| child 61799 | 4cf66f21b764 |
| permissions | -rw-r--r-- |
| 58882 | 1 |
section \<open>Using Hoare Logic\<close> |
| 10148 | 2 |
|
| 31758 | 3 |
theory Hoare_Ex |
4 |
imports Hoare |
|
5 |
begin |
|
| 10148 | 6 |
|
| 58614 | 7 |
subsection \<open>State spaces\<close> |
| 10148 | 8 |
|
| 61541 | 9 |
text \<open>First of all we provide a store of program variables that occur in any |
10 |
of the programs considered later. Slightly unexpected things may happen |
|
11 |
when attempting to work with undeclared variables.\<close> |
|
| 10148 | 12 |
|
13 |
record vars = |
|
14 |
I :: nat |
|
15 |
M :: nat |
|
16 |
N :: nat |
|
17 |
S :: nat |
|
18 |
||
| 61541 | 19 |
text \<open>While all of our variables happen to have the same type, nothing would |
20 |
prevent us from working with many-sorted programs as well, or even |
|
21 |
polymorphic ones. Also note that Isabelle/HOL's extensible record types |
|
22 |
even provides simple means to extend the state space later.\<close> |
|
| 10148 | 23 |
|
24 |
||
| 58614 | 25 |
subsection \<open>Basic examples\<close> |
| 10148 | 26 |
|
| 61541 | 27 |
text \<open>We look at few trivialities involving assignment and sequential |
28 |
composition, in order to get an idea of how to work with our formulation |
|
29 |
of Hoare Logic.\<close> |
|
| 10148 | 30 |
|
| 61541 | 31 |
text \<open>Using the basic \<open>assign\<close> rule directly is a bit |
| 58614 | 32 |
cumbersome.\<close> |
| 10148 | 33 |
|
| 55656 | 34 |
lemma "\<turnstile> \<lbrace>\<acute>(N_update (\<lambda>_. (2 * \<acute>N))) \<in> \<lbrace>\<acute>N = 10\<rbrace>\<rbrace> \<acute>N := 2 * \<acute>N \<lbrace>\<acute>N = 10\<rbrace>" |
| 10148 | 35 |
by (rule assign) |
36 |
||
| 61541 | 37 |
text \<open>Certainly we want the state modification already done, e.g.\ by |
38 |
simplification. The \<open>hoare\<close> method performs the basic state update for us; |
|
39 |
we may apply the Simplifier afterwards to achieve ``obvious'' consequences |
|
40 |
as well.\<close> |
|
| 10148 | 41 |
|
| 55656 | 42 |
lemma "\<turnstile> \<lbrace>True\<rbrace> \<acute>N := 10 \<lbrace>\<acute>N = 10\<rbrace>" |
| 10148 | 43 |
by hoare |
44 |
||
| 55656 | 45 |
lemma "\<turnstile> \<lbrace>2 * \<acute>N = 10\<rbrace> \<acute>N := 2 * \<acute>N \<lbrace>\<acute>N = 10\<rbrace>" |
| 10148 | 46 |
by hoare |
47 |
||
| 55656 | 48 |
lemma "\<turnstile> \<lbrace>\<acute>N = 5\<rbrace> \<acute>N := 2 * \<acute>N \<lbrace>\<acute>N = 10\<rbrace>" |
| 10148 | 49 |
by hoare simp |
50 |
||
| 55656 | 51 |
lemma "\<turnstile> \<lbrace>\<acute>N + 1 = a + 1\<rbrace> \<acute>N := \<acute>N + 1 \<lbrace>\<acute>N = a + 1\<rbrace>" |
| 10148 | 52 |
by hoare |
53 |
||
| 55656 | 54 |
lemma "\<turnstile> \<lbrace>\<acute>N = a\<rbrace> \<acute>N := \<acute>N + 1 \<lbrace>\<acute>N = a + 1\<rbrace>" |
| 10148 | 55 |
by hoare simp |
56 |
||
| 55656 | 57 |
lemma "\<turnstile> \<lbrace>a = a \<and> b = b\<rbrace> \<acute>M := a; \<acute>N := b \<lbrace>\<acute>M = a \<and> \<acute>N = b\<rbrace>" |
| 10148 | 58 |
by hoare |
59 |
||
| 55656 | 60 |
lemma "\<turnstile> \<lbrace>True\<rbrace> \<acute>M := a; \<acute>N := b \<lbrace>\<acute>M = a \<and> \<acute>N = b\<rbrace>" |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55656
diff
changeset
|
61 |
by hoare |
| 10148 | 62 |
|
63 |
lemma |
|
| 55656 | 64 |
"\<turnstile> \<lbrace>\<acute>M = a \<and> \<acute>N = b\<rbrace> |
| 46582 | 65 |
\<acute>I := \<acute>M; \<acute>M := \<acute>N; \<acute>N := \<acute>I |
| 55656 | 66 |
\<lbrace>\<acute>M = b \<and> \<acute>N = a\<rbrace>" |
| 10148 | 67 |
by hoare simp |
68 |
||
| 61541 | 69 |
text \<open>It is important to note that statements like the following one can |
70 |
only be proven for each individual program variable. Due to the |
|
| 37671 | 71 |
extra-logical nature of record fields, we cannot formulate a theorem |
| 58614 | 72 |
relating record selectors and updates schematically.\<close> |
| 10148 | 73 |
|
| 55656 | 74 |
lemma "\<turnstile> \<lbrace>\<acute>N = a\<rbrace> \<acute>N := \<acute>N \<lbrace>\<acute>N = a\<rbrace>" |
| 10148 | 75 |
by hoare |
76 |
||
| 55656 | 77 |
lemma "\<turnstile> \<lbrace>\<acute>x = a\<rbrace> \<acute>x := \<acute>x \<lbrace>\<acute>x = a\<rbrace>" |
| 10148 | 78 |
oops |
79 |
||
80 |
lemma |
|
81 |
"Valid {s. x s = a} (Basic (\<lambda>s. x_update (x s) s)) {s. x s = n}"
|
|
| 58614 | 82 |
-- \<open>same statement without concrete syntax\<close> |
| 10148 | 83 |
oops |
84 |
||
85 |
||
| 61541 | 86 |
text \<open>In the following assignments we make use of the consequence rule in |
87 |
order to achieve the intended precondition. Certainly, the \<open>hoare\<close> method |
|
88 |
is able to handle this case, too.\<close> |
|
| 10148 | 89 |
|
| 55656 | 90 |
lemma "\<turnstile> \<lbrace>\<acute>M = \<acute>N\<rbrace> \<acute>M := \<acute>M + 1 \<lbrace>\<acute>M \<noteq> \<acute>N\<rbrace>" |
| 10148 | 91 |
proof - |
| 55656 | 92 |
have "\<lbrace>\<acute>M = \<acute>N\<rbrace> \<subseteq> \<lbrace>\<acute>M + 1 \<noteq> \<acute>N\<rbrace>" |
| 10148 | 93 |
by auto |
| 55656 | 94 |
also have "\<turnstile> \<dots> \<acute>M := \<acute>M + 1 \<lbrace>\<acute>M \<noteq> \<acute>N\<rbrace>" |
| 10148 | 95 |
by hoare |
96 |
finally show ?thesis . |
|
97 |
qed |
|
98 |
||
| 55656 | 99 |
lemma "\<turnstile> \<lbrace>\<acute>M = \<acute>N\<rbrace> \<acute>M := \<acute>M + 1 \<lbrace>\<acute>M \<noteq> \<acute>N\<rbrace>" |
| 10148 | 100 |
proof - |
| 60410 | 101 |
have "m = n \<longrightarrow> m + 1 \<noteq> n" for m n :: nat |
| 58614 | 102 |
-- \<open>inclusion of assertions expressed in ``pure'' logic,\<close> |
103 |
-- \<open>without mentioning the state space\<close> |
|
| 10148 | 104 |
by simp |
| 55656 | 105 |
also have "\<turnstile> \<lbrace>\<acute>M + 1 \<noteq> \<acute>N\<rbrace> \<acute>M := \<acute>M + 1 \<lbrace>\<acute>M \<noteq> \<acute>N\<rbrace>" |
| 10148 | 106 |
by hoare |
107 |
finally show ?thesis . |
|
108 |
qed |
|
109 |
||
| 55656 | 110 |
lemma "\<turnstile> \<lbrace>\<acute>M = \<acute>N\<rbrace> \<acute>M := \<acute>M + 1 \<lbrace>\<acute>M \<noteq> \<acute>N\<rbrace>" |
| 10148 | 111 |
by hoare simp |
112 |
||
113 |
||
| 58614 | 114 |
subsection \<open>Multiplication by addition\<close> |
| 10148 | 115 |
|
| 61541 | 116 |
text \<open>We now do some basic examples of actual \<^verbatim>\<open>WHILE\<close> programs. This one is |
117 |
a loop for calculating the product of two natural numbers, by iterated |
|
118 |
addition. We first give detailed structured proof based on single-step |
|
119 |
Hoare rules.\<close> |
|
| 10148 | 120 |
|
121 |
lemma |
|
| 55656 | 122 |
"\<turnstile> \<lbrace>\<acute>M = 0 \<and> \<acute>S = 0\<rbrace> |
123 |
WHILE \<acute>M \<noteq> a |
|
| 10838 | 124 |
DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD |
| 55656 | 125 |
\<lbrace>\<acute>S = a * b\<rbrace>" |
| 10148 | 126 |
proof - |
| 55656 | 127 |
let "\<turnstile> _ ?while _" = ?thesis |
128 |
let "\<lbrace>\<acute>?inv\<rbrace>" = "\<lbrace>\<acute>S = \<acute>M * b\<rbrace>" |
|
| 10148 | 129 |
|
| 55656 | 130 |
have "\<lbrace>\<acute>M = 0 \<and> \<acute>S = 0\<rbrace> \<subseteq> \<lbrace>\<acute>?inv\<rbrace>" by auto |
131 |
also have "\<turnstile> \<dots> ?while \<lbrace>\<acute>?inv \<and> \<not> (\<acute>M \<noteq> a)\<rbrace>" |
|
| 10148 | 132 |
proof |
| 10838 | 133 |
let ?c = "\<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1" |
| 55656 | 134 |
have "\<lbrace>\<acute>?inv \<and> \<acute>M \<noteq> a\<rbrace> \<subseteq> \<lbrace>\<acute>S + b = (\<acute>M + 1) * b\<rbrace>" |
| 10148 | 135 |
by auto |
| 55656 | 136 |
also have "\<turnstile> \<dots> ?c \<lbrace>\<acute>?inv\<rbrace>" by hoare |
137 |
finally show "\<turnstile> \<lbrace>\<acute>?inv \<and> \<acute>M \<noteq> a\<rbrace> ?c \<lbrace>\<acute>?inv\<rbrace>" . |
|
| 10148 | 138 |
qed |
| 55656 | 139 |
also have "\<dots> \<subseteq> \<lbrace>\<acute>S = a * b\<rbrace>" by auto |
| 10148 | 140 |
finally show ?thesis . |
141 |
qed |
|
142 |
||
| 61541 | 143 |
text \<open>The subsequent version of the proof applies the \<open>hoare\<close> method to |
144 |
reduce the Hoare statement to a purely logical problem that can be solved |
|
145 |
fully automatically. Note that we have to specify the \<^verbatim>\<open>WHILE\<close> loop |
|
146 |
invariant in the original statement.\<close> |
|
| 10148 | 147 |
|
148 |
lemma |
|
| 55656 | 149 |
"\<turnstile> \<lbrace>\<acute>M = 0 \<and> \<acute>S = 0\<rbrace> |
150 |
WHILE \<acute>M \<noteq> a |
|
151 |
INV \<lbrace>\<acute>S = \<acute>M * b\<rbrace> |
|
| 10838 | 152 |
DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD |
| 55656 | 153 |
\<lbrace>\<acute>S = a * b\<rbrace>" |
| 10148 | 154 |
by hoare auto |
155 |
||
156 |
||
| 58614 | 157 |
subsection \<open>Summing natural numbers\<close> |
| 10148 | 158 |
|
| 61541 | 159 |
text \<open>We verify an imperative program to sum natural numbers up to a given |
160 |
limit. First some functional definition for proper specification of the |
|
161 |
problem. |
|
| 10148 | 162 |
|
| 61541 | 163 |
\<^medskip> |
164 |
The following proof is quite explicit in the individual steps taken, with |
|
165 |
the \<open>hoare\<close> method only applied locally to take care of assignment and |
|
166 |
sequential composition. Note that we express intermediate proof obligation |
|
167 |
in pure logic, without referring to the state space.\<close> |
|
| 15569 | 168 |
|
| 10148 | 169 |
theorem |
| 55656 | 170 |
"\<turnstile> \<lbrace>True\<rbrace> |
| 10838 | 171 |
\<acute>S := 0; \<acute>I := 1; |
| 55656 | 172 |
WHILE \<acute>I \<noteq> n |
| 10148 | 173 |
DO |
| 10838 | 174 |
\<acute>S := \<acute>S + \<acute>I; |
175 |
\<acute>I := \<acute>I + 1 |
|
| 10148 | 176 |
OD |
| 55656 | 177 |
\<lbrace>\<acute>S = (\<Sum>j<n. j)\<rbrace>" |
178 |
(is "\<turnstile> _ (_; ?while) _") |
|
| 10148 | 179 |
proof - |
| 55656 | 180 |
let ?sum = "\<lambda>k::nat. \<Sum>j<k. j" |
| 15049 | 181 |
let ?inv = "\<lambda>s i::nat. s = ?sum i" |
| 10148 | 182 |
|
| 55656 | 183 |
have "\<turnstile> \<lbrace>True\<rbrace> \<acute>S := 0; \<acute>I := 1 \<lbrace>?inv \<acute>S \<acute>I\<rbrace>" |
| 10148 | 184 |
proof - |
| 55656 | 185 |
have "True \<longrightarrow> 0 = ?sum 1" |
| 10148 | 186 |
by simp |
| 55656 | 187 |
also have "\<turnstile> \<lbrace>\<dots>\<rbrace> \<acute>S := 0; \<acute>I := 1 \<lbrace>?inv \<acute>S \<acute>I\<rbrace>" |
| 10148 | 188 |
by hoare |
189 |
finally show ?thesis . |
|
190 |
qed |
|
| 55656 | 191 |
also have "\<turnstile> \<dots> ?while \<lbrace>?inv \<acute>S \<acute>I \<and> \<not> \<acute>I \<noteq> n\<rbrace>" |
| 10148 | 192 |
proof |
| 10838 | 193 |
let ?body = "\<acute>S := \<acute>S + \<acute>I; \<acute>I := \<acute>I + 1" |
| 60410 | 194 |
have "?inv s i \<and> i \<noteq> n \<longrightarrow> ?inv (s + i) (i + 1)" for s i |
| 10148 | 195 |
by simp |
| 55656 | 196 |
also have "\<turnstile> \<lbrace>\<acute>S + \<acute>I = ?sum (\<acute>I + 1)\<rbrace> ?body \<lbrace>?inv \<acute>S \<acute>I\<rbrace>" |
| 10148 | 197 |
by hoare |
| 55656 | 198 |
finally show "\<turnstile> \<lbrace>?inv \<acute>S \<acute>I \<and> \<acute>I \<noteq> n\<rbrace> ?body \<lbrace>?inv \<acute>S \<acute>I\<rbrace>" . |
| 10148 | 199 |
qed |
| 60410 | 200 |
also have "s = ?sum i \<and> \<not> i \<noteq> n \<longrightarrow> s = ?sum n" for s i |
| 10148 | 201 |
by simp |
202 |
finally show ?thesis . |
|
203 |
qed |
|
204 |
||
| 61541 | 205 |
text \<open>The next version uses the \<open>hoare\<close> method, while still explaining the |
206 |
resulting proof obligations in an abstract, structured manner.\<close> |
|
| 10148 | 207 |
|
208 |
theorem |
|
| 55656 | 209 |
"\<turnstile> \<lbrace>True\<rbrace> |
| 10838 | 210 |
\<acute>S := 0; \<acute>I := 1; |
| 55656 | 211 |
WHILE \<acute>I \<noteq> n |
212 |
INV \<lbrace>\<acute>S = (\<Sum>j<\<acute>I. j)\<rbrace> |
|
| 10148 | 213 |
DO |
| 10838 | 214 |
\<acute>S := \<acute>S + \<acute>I; |
215 |
\<acute>I := \<acute>I + 1 |
|
| 10148 | 216 |
OD |
| 55656 | 217 |
\<lbrace>\<acute>S = (\<Sum>j<n. j)\<rbrace>" |
| 10148 | 218 |
proof - |
| 55656 | 219 |
let ?sum = "\<lambda>k::nat. \<Sum>j<k. j" |
| 15049 | 220 |
let ?inv = "\<lambda>s i::nat. s = ?sum i" |
| 10148 | 221 |
show ?thesis |
222 |
proof hoare |
|
223 |
show "?inv 0 1" by simp |
|
| 60416 | 224 |
show "?inv (s + i) (i + 1)" if "?inv s i \<and> i \<noteq> n" for s i |
| 60449 | 225 |
using that by simp |
| 60416 | 226 |
show "s = ?sum n" if "?inv s i \<and> \<not> i \<noteq> n" for s i |
| 60449 | 227 |
using that by simp |
| 10148 | 228 |
qed |
229 |
qed |
|
230 |
||
| 61541 | 231 |
text \<open>Certainly, this proof may be done fully automatic as well, provided |
232 |
that the invariant is given beforehand.\<close> |
|
| 10148 | 233 |
|
234 |
theorem |
|
| 55656 | 235 |
"\<turnstile> \<lbrace>True\<rbrace> |
| 10838 | 236 |
\<acute>S := 0; \<acute>I := 1; |
| 55656 | 237 |
WHILE \<acute>I \<noteq> n |
238 |
INV \<lbrace>\<acute>S = (\<Sum>j<\<acute>I. j)\<rbrace> |
|
| 10148 | 239 |
DO |
| 10838 | 240 |
\<acute>S := \<acute>S + \<acute>I; |
241 |
\<acute>I := \<acute>I + 1 |
|
| 10148 | 242 |
OD |
| 55656 | 243 |
\<lbrace>\<acute>S = (\<Sum>j<n. j)\<rbrace>" |
| 10148 | 244 |
by hoare auto |
245 |
||
| 18193 | 246 |
|
| 58614 | 247 |
subsection \<open>Time\<close> |
| 13473 | 248 |
|
| 61541 | 249 |
text \<open>A simple embedding of time in Hoare logic: function \<open>timeit\<close> inserts |
250 |
an extra variable to keep track of the elapsed time.\<close> |
|
| 13473 | 251 |
|
252 |
record tstate = time :: nat |
|
253 |
||
| 41818 | 254 |
type_synonym 'a time = "\<lparr>time :: nat, \<dots> :: 'a\<rparr>" |
| 13473 | 255 |
|
| 37671 | 256 |
primrec timeit :: "'a time com \<Rightarrow> 'a time com" |
257 |
where |
|
| 18193 | 258 |
"timeit (Basic f) = (Basic f; Basic(\<lambda>s. s\<lparr>time := Suc (time s)\<rparr>))" |
| 37671 | 259 |
| "timeit (c1; c2) = (timeit c1; timeit c2)" |
260 |
| "timeit (Cond b c1 c2) = Cond b (timeit c1) (timeit c2)" |
|
261 |
| "timeit (While b iv c) = While b iv (timeit c)" |
|
| 13473 | 262 |
|
263 |
record tvars = tstate + |
|
264 |
I :: nat |
|
265 |
J :: nat |
|
266 |
||
| 18193 | 267 |
lemma lem: "(0::nat) < n \<Longrightarrow> n + n \<le> Suc (n * n)" |
268 |
by (induct n) simp_all |
|
| 13473 | 269 |
|
| 46582 | 270 |
lemma |
| 55656 | 271 |
"\<turnstile> \<lbrace>i = \<acute>I \<and> \<acute>time = 0\<rbrace> |
272 |
(timeit |
|
273 |
(WHILE \<acute>I \<noteq> 0 |
|
274 |
INV \<lbrace>2 *\<acute> time + \<acute>I * \<acute>I + 5 * \<acute>I = i * i + 5 * i\<rbrace> |
|
275 |
DO |
|
276 |
\<acute>J := \<acute>I; |
|
277 |
WHILE \<acute>J \<noteq> 0 |
|
278 |
INV \<lbrace>0 < \<acute>I \<and> 2 * \<acute>time + \<acute>I * \<acute>I + 3 * \<acute>I + 2 * \<acute>J - 2 = i * i + 5 * i\<rbrace> |
|
279 |
DO \<acute>J := \<acute>J - 1 OD; |
|
280 |
\<acute>I := \<acute>I - 1 |
|
281 |
OD)) |
|
282 |
\<lbrace>2 * \<acute>time = i * i + 5 * i\<rbrace>" |
|
| 18193 | 283 |
apply simp |
284 |
apply hoare |
|
285 |
apply simp |
|
286 |
apply clarsimp |
|
287 |
apply clarsimp |
|
|
20432
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
webertj
parents:
20272
diff
changeset
|
288 |
apply arith |
| 18193 | 289 |
prefer 2 |
| 13473 | 290 |
apply clarsimp |
| 18193 | 291 |
apply (clarsimp simp: nat_distrib) |
292 |
apply (frule lem) |
|
| 13473 | 293 |
apply arith |
| 18193 | 294 |
done |
| 13473 | 295 |
|
296 |
end |