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