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