| author | wenzelm | 
| Fri, 16 Jan 2009 15:21:26 +0100 | |
| changeset 29517 | d7648f30f923 | 
| parent 26936 | faf8a5b5ba87 | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 1479 | 1 | (* Title: HOLCF/ex/hoare.thy | 
| 244 | 2 | ID: $Id$ | 
| 1479 | 3 | Author: Franz Regensburger | 
| 244 | 4 | |
| 12036 | 5 | Theory for an example by C.A.R. Hoare | 
| 244 | 6 | |
| 17291 | 7 | p x = if b1 x | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
1150diff
changeset | 8 | then p (g x) | 
| 244 | 9 | else x fi | 
| 10 | ||
| 17291 | 11 | q x = if b1 x orelse b2 x | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
1150diff
changeset | 12 | then q (g x) | 
| 244 | 13 | else x fi | 
| 14 | ||
| 17291 | 15 | Prove: for all b1 b2 g . | 
| 16 | q o p = q | |
| 244 | 17 | |
| 18 | In order to get a nice notation we fix the functions b1,b2 and g in the | |
| 19 | signature of this example | |
| 20 | ||
| 21 | *) | |
| 22 | ||
| 17291 | 23 | theory Hoare | 
| 24 | imports HOLCF | |
| 25 | begin | |
| 244 | 26 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
21404diff
changeset | 27 | axiomatization | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
21404diff
changeset | 28 | b1 :: "'a -> tr" and | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
21404diff
changeset | 29 | b2 :: "'a -> tr" and | 
| 17291 | 30 | g :: "'a -> 'a" | 
| 244 | 31 | |
| 19763 | 32 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 33 | p :: "'a -> 'a" where | 
| 19763 | 34 | "p = fix$(LAM f. LAM x. If b1$x then f$(g$x) else x fi)" | 
| 244 | 35 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 36 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 37 | q :: "'a -> 'a" where | 
| 19763 | 38 | "q = fix$(LAM f. LAM x. If b1$x orelse b2$x then f$(g$x) else x fi)" | 
| 244 | 39 | |
| 19742 | 40 | |
| 41 | (* --------- pure HOLCF logic, some little lemmas ------ *) | |
| 42 | ||
| 43 | lemma hoare_lemma2: "b~=TT ==> b=FF | b=UU" | |
| 44 | apply (rule Exh_tr [THEN disjE]) | |
| 45 | apply blast+ | |
| 46 | done | |
| 47 | ||
| 48 | lemma hoare_lemma3: " (ALL k. b1$(iterate k$g$x) = TT) | (EX k. b1$(iterate k$g$x)~=TT)" | |
| 49 | apply blast | |
| 50 | done | |
| 51 | ||
| 52 | lemma hoare_lemma4: "(EX k. b1$(iterate k$g$x) ~= TT) ==> | |
| 53 | EX k. b1$(iterate k$g$x) = FF | b1$(iterate k$g$x) = UU" | |
| 54 | apply (erule exE) | |
| 55 | apply (rule exI) | |
| 56 | apply (rule hoare_lemma2) | |
| 57 | apply assumption | |
| 58 | done | |
| 59 | ||
| 60 | lemma hoare_lemma5: "[|(EX k. b1$(iterate k$g$x) ~= TT); | |
| 61 | k=Least(%n. b1$(iterate n$g$x) ~= TT)|] ==> | |
| 62 | b1$(iterate k$g$x)=FF | b1$(iterate k$g$x)=UU" | |
| 63 | apply hypsubst | |
| 64 | apply (rule hoare_lemma2) | |
| 65 | apply (erule exE) | |
| 66 | apply (erule LeastI) | |
| 67 | done | |
| 68 | ||
| 69 | lemma hoare_lemma6: "b=UU ==> b~=TT" | |
| 70 | apply hypsubst | |
| 71 | apply (rule dist_eq_tr) | |
| 72 | done | |
| 73 | ||
| 74 | lemma hoare_lemma7: "b=FF ==> b~=TT" | |
| 75 | apply hypsubst | |
| 76 | apply (rule dist_eq_tr) | |
| 77 | done | |
| 78 | ||
| 79 | lemma hoare_lemma8: "[|(EX k. b1$(iterate k$g$x) ~= TT); | |
| 80 | k=Least(%n. b1$(iterate n$g$x) ~= TT)|] ==> | |
| 81 | ALL m. m < k --> b1$(iterate m$g$x)=TT" | |
| 82 | apply hypsubst | |
| 83 | apply (erule exE) | |
| 84 | apply (intro strip) | |
| 85 | apply (rule_tac p = "b1$ (iterate m$g$x) " in trE) | |
| 86 | prefer 2 apply (assumption) | |
| 26334 | 87 | apply (rule le_less_trans [THEN less_irrefl [THEN notE]]) | 
| 19742 | 88 | prefer 2 apply (assumption) | 
| 89 | apply (rule Least_le) | |
| 90 | apply (erule hoare_lemma6) | |
| 26334 | 91 | apply (rule le_less_trans [THEN less_irrefl [THEN notE]]) | 
| 19742 | 92 | prefer 2 apply (assumption) | 
| 93 | apply (rule Least_le) | |
| 94 | apply (erule hoare_lemma7) | |
| 95 | done | |
| 96 | ||
| 97 | ||
| 98 | lemma hoare_lemma28: "f$(y::'a)=(UU::tr) ==> f$UU = UU" | |
| 99 | apply (erule flat_codom [THEN disjE]) | |
| 100 | apply auto | |
| 101 | done | |
| 102 | ||
| 103 | ||
| 104 | (* ----- access to definitions ----- *) | |
| 105 | ||
| 106 | lemma p_def3: "p$x = If b1$x then p$(g$x) else x fi" | |
| 107 | apply (rule trans) | |
| 19763 | 108 | apply (rule p_def [THEN eq_reflection, THEN fix_eq3]) | 
| 19742 | 109 | apply simp | 
| 110 | done | |
| 111 | ||
| 112 | lemma q_def3: "q$x = If b1$x orelse b2$x then q$(g$x) else x fi" | |
| 113 | apply (rule trans) | |
| 19763 | 114 | apply (rule q_def [THEN eq_reflection, THEN fix_eq3]) | 
| 19742 | 115 | apply simp | 
| 116 | done | |
| 117 | ||
| 118 | (** --------- proofs about iterations of p and q ---------- **) | |
| 119 | ||
| 120 | lemma hoare_lemma9: "(ALL m. m< Suc k --> b1$(iterate m$g$x)=TT) --> | |
| 121 | p$(iterate k$g$x)=p$x" | |
| 122 | apply (induct_tac k) | |
| 123 | apply (simp (no_asm)) | |
| 124 | apply (simp (no_asm)) | |
| 125 | apply (intro strip) | |
| 126 | apply (rule_tac s = "p$ (iterate n$g$x) " in trans) | |
| 127 | apply (rule trans) | |
| 128 | apply (rule_tac [2] p_def3 [symmetric]) | |
| 129 | apply (rule_tac s = "TT" and t = "b1$ (iterate n$g$x) " in ssubst) | |
| 130 | apply (rule mp) | |
| 131 | apply (erule spec) | |
| 132 | apply (simp (no_asm) add: less_Suc_eq) | |
| 133 | apply simp | |
| 134 | apply (erule mp) | |
| 135 | apply (intro strip) | |
| 136 | apply (rule mp) | |
| 137 | apply (erule spec) | |
| 138 | apply (erule less_trans) | |
| 139 | apply simp | |
| 140 | done | |
| 141 | ||
| 142 | lemma hoare_lemma24: "(ALL m. m< Suc k --> b1$(iterate m$g$x)=TT) --> | |
| 143 | q$(iterate k$g$x)=q$x" | |
| 144 | apply (induct_tac k) | |
| 145 | apply (simp (no_asm)) | |
| 146 | apply (simp (no_asm) add: less_Suc_eq) | |
| 147 | apply (intro strip) | |
| 148 | apply (rule_tac s = "q$ (iterate n$g$x) " in trans) | |
| 149 | apply (rule trans) | |
| 150 | apply (rule_tac [2] q_def3 [symmetric]) | |
| 151 | apply (rule_tac s = "TT" and t = "b1$ (iterate n$g$x) " in ssubst) | |
| 152 | apply blast | |
| 153 | apply simp | |
| 154 | apply (erule mp) | |
| 155 | apply (intro strip) | |
| 156 | apply (fast dest!: less_Suc_eq [THEN iffD1]) | |
| 157 | done | |
| 158 | ||
| 159 | (* -------- results about p for case (EX k. b1$(iterate k$g$x)~=TT) ------- *) | |
| 160 | ||
| 26936 | 161 | thm hoare_lemma8 [THEN hoare_lemma9 [THEN mp], standard] | 
| 19742 | 162 | |
| 163 | lemma hoare_lemma10: | |
| 164 | "EX k. b1$(iterate k$g$x) ~= TT | |
| 165 | ==> Suc k = (LEAST n. b1$(iterate n$g$x) ~= TT) ==> p$(iterate k$g$x) = p$x" | |
| 166 | by (rule hoare_lemma8 [THEN hoare_lemma9 [THEN mp]]) | |
| 167 | ||
| 168 | lemma hoare_lemma11: "(EX n. b1$(iterate n$g$x) ~= TT) ==> | |
| 169 | k=(LEAST n. b1$(iterate n$g$x) ~= TT) & b1$(iterate k$g$x)=FF | |
| 170 | --> p$x = iterate k$g$x" | |
| 171 | apply (case_tac "k") | |
| 172 | apply hypsubst | |
| 173 | apply (simp (no_asm)) | |
| 174 | apply (intro strip) | |
| 175 | apply (erule conjE) | |
| 176 | apply (rule trans) | |
| 177 | apply (rule p_def3) | |
| 178 | apply simp | |
| 179 | apply hypsubst | |
| 180 | apply (intro strip) | |
| 181 | apply (erule conjE) | |
| 182 | apply (rule trans) | |
| 183 | apply (erule hoare_lemma10 [symmetric]) | |
| 184 | apply assumption | |
| 185 | apply (rule trans) | |
| 186 | apply (rule p_def3) | |
| 187 | apply (rule_tac s = "TT" and t = "b1$ (iterate nat$g$x) " in ssubst) | |
| 188 | apply (rule hoare_lemma8 [THEN spec, THEN mp]) | |
| 189 | apply assumption | |
| 190 | apply assumption | |
| 191 | apply (simp (no_asm)) | |
| 192 | apply (tactic "simp_tac HOLCF_ss 1") | |
| 193 | apply (rule trans) | |
| 194 | apply (rule p_def3) | |
| 195 | apply (simp (no_asm) del: iterate_Suc add: iterate_Suc [symmetric]) | |
| 196 | apply (erule_tac s = "FF" in ssubst) | |
| 197 | apply simp | |
| 198 | done | |
| 199 | ||
| 200 | lemma hoare_lemma12: "(EX n. b1$(iterate n$g$x) ~= TT) ==> | |
| 201 | k=Least(%n. b1$(iterate n$g$x)~=TT) & b1$(iterate k$g$x)=UU | |
| 202 | --> p$x = UU" | |
| 203 | apply (case_tac "k") | |
| 204 | apply hypsubst | |
| 205 | apply (simp (no_asm)) | |
| 206 | apply (intro strip) | |
| 207 | apply (erule conjE) | |
| 208 | apply (rule trans) | |
| 209 | apply (rule p_def3) | |
| 210 | apply simp | |
| 211 | apply hypsubst | |
| 212 | apply (simp (no_asm)) | |
| 213 | apply (intro strip) | |
| 214 | apply (erule conjE) | |
| 215 | apply (rule trans) | |
| 216 | apply (rule hoare_lemma10 [symmetric]) | |
| 217 | apply assumption | |
| 218 | apply assumption | |
| 219 | apply (rule trans) | |
| 220 | apply (rule p_def3) | |
| 221 | apply (rule_tac s = "TT" and t = "b1$ (iterate nat$g$x) " in ssubst) | |
| 222 | apply (rule hoare_lemma8 [THEN spec, THEN mp]) | |
| 223 | apply assumption | |
| 224 | apply assumption | |
| 225 | apply (simp (no_asm)) | |
| 226 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 227 | apply (rule trans) | |
| 228 | apply (rule p_def3) | |
| 229 | apply simp | |
| 230 | done | |
| 231 | ||
| 232 | (* -------- results about p for case (ALL k. b1$(iterate k$g$x)=TT) ------- *) | |
| 233 | ||
| 234 | lemma fernpass_lemma: "(ALL k. b1$(iterate k$g$x)=TT) ==> ALL k. p$(iterate k$g$x) = UU" | |
| 19763 | 235 | apply (rule p_def [THEN eq_reflection, THEN def_fix_ind]) | 
| 19742 | 236 | apply (rule adm_all) | 
| 237 | apply (rule adm_eq) | |
| 238 | apply (tactic "cont_tacR 1") | |
| 239 | apply (rule allI) | |
| 240 | apply (subst Rep_CFun_strict1) | |
| 241 | apply (rule refl) | |
| 242 | apply (simp (no_asm)) | |
| 243 | apply (rule allI) | |
| 244 | apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$x) " in ssubst) | |
| 245 | apply (erule spec) | |
| 246 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 247 | apply (rule iterate_Suc [THEN subst]) | |
| 248 | apply (erule spec) | |
| 249 | done | |
| 250 | ||
| 251 | lemma hoare_lemma16: "(ALL k. b1$(iterate k$g$x)=TT) ==> p$x = UU" | |
| 252 | apply (rule_tac F1 = "g" and t = "x" in iterate_0 [THEN subst]) | |
| 253 | apply (erule fernpass_lemma [THEN spec]) | |
| 254 | done | |
| 255 | ||
| 256 | (* -------- results about q for case (ALL k. b1$(iterate k$g$x)=TT) ------- *) | |
| 257 | ||
| 258 | lemma hoare_lemma17: "(ALL k. b1$(iterate k$g$x)=TT) ==> ALL k. q$(iterate k$g$x) = UU" | |
| 19763 | 259 | apply (rule q_def [THEN eq_reflection, THEN def_fix_ind]) | 
| 19742 | 260 | apply (rule adm_all) | 
| 261 | apply (rule adm_eq) | |
| 262 | apply (tactic "cont_tacR 1") | |
| 263 | apply (rule allI) | |
| 264 | apply (subst Rep_CFun_strict1) | |
| 265 | apply (rule refl) | |
| 266 | apply (rule allI) | |
| 267 | apply (simp (no_asm)) | |
| 268 | apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$x) " in ssubst) | |
| 269 | apply (erule spec) | |
| 270 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 271 | apply (rule iterate_Suc [THEN subst]) | |
| 272 | apply (erule spec) | |
| 273 | done | |
| 274 | ||
| 275 | lemma hoare_lemma18: "(ALL k. b1$(iterate k$g$x)=TT) ==> q$x = UU" | |
| 276 | apply (rule_tac F1 = "g" and t = "x" in iterate_0 [THEN subst]) | |
| 277 | apply (erule hoare_lemma17 [THEN spec]) | |
| 278 | done | |
| 279 | ||
| 280 | lemma hoare_lemma19: | |
| 281 | "(ALL k. (b1::'a->tr)$(iterate k$g$x)=TT) ==> b1$(UU::'a) = UU | (ALL y. b1$(y::'a)=TT)" | |
| 282 | apply (rule flat_codom) | |
| 283 | apply (rule_tac t = "x1" in iterate_0 [THEN subst]) | |
| 284 | apply (erule spec) | |
| 285 | done | |
| 286 | ||
| 287 | lemma hoare_lemma20: "(ALL y. b1$(y::'a)=TT) ==> ALL k. q$(iterate k$g$(x::'a)) = UU" | |
| 19763 | 288 | apply (rule q_def [THEN eq_reflection, THEN def_fix_ind]) | 
| 19742 | 289 | apply (rule adm_all) | 
| 290 | apply (rule adm_eq) | |
| 291 | apply (tactic "cont_tacR 1") | |
| 292 | apply (rule allI) | |
| 293 | apply (subst Rep_CFun_strict1) | |
| 294 | apply (rule refl) | |
| 295 | apply (rule allI) | |
| 296 | apply (simp (no_asm)) | |
| 297 | apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$ (x::'a))" in ssubst) | |
| 298 | apply (erule spec) | |
| 299 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 300 | apply (rule iterate_Suc [THEN subst]) | |
| 301 | apply (erule spec) | |
| 302 | done | |
| 303 | ||
| 304 | lemma hoare_lemma21: "(ALL y. b1$(y::'a)=TT) ==> q$(x::'a) = UU" | |
| 305 | apply (rule_tac F1 = "g" and t = "x" in iterate_0 [THEN subst]) | |
| 306 | apply (erule hoare_lemma20 [THEN spec]) | |
| 307 | done | |
| 308 | ||
| 309 | lemma hoare_lemma22: "b1$(UU::'a)=UU ==> q$(UU::'a) = UU" | |
| 310 | apply (subst q_def3) | |
| 311 | apply simp | |
| 312 | done | |
| 313 | ||
| 314 | (* -------- results about q for case (EX k. b1$(iterate k$g$x) ~= TT) ------- *) | |
| 315 | ||
| 316 | lemma hoare_lemma25: "EX k. b1$(iterate k$g$x) ~= TT | |
| 317 | ==> Suc k = (LEAST n. b1$(iterate n$g$x) ~= TT) ==> q$(iterate k$g$x) = q$x" | |
| 318 | by (rule hoare_lemma8 [THEN hoare_lemma24 [THEN mp]]) | |
| 319 | ||
| 320 | lemma hoare_lemma26: "(EX n. b1$(iterate n$g$x)~=TT) ==> | |
| 321 | k=Least(%n. b1$(iterate n$g$x) ~= TT) & b1$(iterate k$g$x) =FF | |
| 322 | --> q$x = q$(iterate k$g$x)" | |
| 323 | apply (case_tac "k") | |
| 324 | apply hypsubst | |
| 325 | apply (intro strip) | |
| 326 | apply (simp (no_asm)) | |
| 327 | apply hypsubst | |
| 328 | apply (intro strip) | |
| 329 | apply (erule conjE) | |
| 330 | apply (rule trans) | |
| 331 | apply (rule hoare_lemma25 [symmetric]) | |
| 332 | apply assumption | |
| 333 | apply assumption | |
| 334 | apply (rule trans) | |
| 335 | apply (rule q_def3) | |
| 336 | apply (rule_tac s = "TT" and t = "b1$ (iterate nat$g$x) " in ssubst) | |
| 337 | apply (rule hoare_lemma8 [THEN spec, THEN mp]) | |
| 338 | apply assumption | |
| 339 | apply assumption | |
| 340 | apply (simp (no_asm)) | |
| 341 | apply (simp (no_asm)) | |
| 342 | done | |
| 343 | ||
| 344 | ||
| 345 | lemma hoare_lemma27: "(EX n. b1$(iterate n$g$x) ~= TT) ==> | |
| 346 | k=Least(%n. b1$(iterate n$g$x)~=TT) & b1$(iterate k$g$x)=UU | |
| 347 | --> q$x = UU" | |
| 348 | apply (case_tac "k") | |
| 349 | apply hypsubst | |
| 350 | apply (simp (no_asm)) | |
| 351 | apply (intro strip) | |
| 352 | apply (erule conjE) | |
| 353 | apply (subst q_def3) | |
| 354 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 355 | apply hypsubst | |
| 356 | apply (simp (no_asm)) | |
| 357 | apply (intro strip) | |
| 358 | apply (erule conjE) | |
| 359 | apply (rule trans) | |
| 360 | apply (rule hoare_lemma25 [symmetric]) | |
| 361 | apply assumption | |
| 362 | apply assumption | |
| 363 | apply (rule trans) | |
| 364 | apply (rule q_def3) | |
| 365 | apply (rule_tac s = "TT" and t = "b1$ (iterate nat$g$x) " in ssubst) | |
| 366 | apply (rule hoare_lemma8 [THEN spec, THEN mp]) | |
| 367 | apply assumption | |
| 368 | apply assumption | |
| 369 | apply (simp (no_asm)) | |
| 370 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 371 | apply (rule trans) | |
| 372 | apply (rule q_def3) | |
| 373 | apply (tactic "asm_simp_tac HOLCF_ss 1") | |
| 374 | done | |
| 375 | ||
| 376 | (* ------- (ALL k. b1$(iterate k$g$x)=TT) ==> q o p = q ----- *) | |
| 377 | ||
| 378 | lemma hoare_lemma23: "(ALL k. b1$(iterate k$g$x)=TT) ==> q$(p$x) = q$x" | |
| 379 | apply (subst hoare_lemma16) | |
| 380 | apply assumption | |
| 381 | apply (rule hoare_lemma19 [THEN disjE]) | |
| 382 | apply assumption | |
| 383 | apply (simplesubst hoare_lemma18) | |
| 384 | apply assumption | |
| 385 | apply (simplesubst hoare_lemma22) | |
| 386 | apply assumption | |
| 387 | apply (rule refl) | |
| 388 | apply (simplesubst hoare_lemma21) | |
| 389 | apply assumption | |
| 390 | apply (simplesubst hoare_lemma21) | |
| 391 | apply assumption | |
| 392 | apply (rule refl) | |
| 393 | done | |
| 394 | ||
| 395 | (* ------------ EX k. b1~(iterate k$g$x) ~= TT ==> q o p = q ----- *) | |
| 396 | ||
| 397 | lemma hoare_lemma29: "EX k. b1$(iterate k$g$x) ~= TT ==> q$(p$x) = q$x" | |
| 398 | apply (rule hoare_lemma5 [THEN disjE]) | |
| 399 | apply assumption | |
| 400 | apply (rule refl) | |
| 401 | apply (subst hoare_lemma11 [THEN mp]) | |
| 402 | apply assumption | |
| 403 | apply (rule conjI) | |
| 404 | apply (rule refl) | |
| 405 | apply assumption | |
| 406 | apply (rule hoare_lemma26 [THEN mp, THEN subst]) | |
| 407 | apply assumption | |
| 408 | apply (rule conjI) | |
| 409 | apply (rule refl) | |
| 410 | apply assumption | |
| 411 | apply (rule refl) | |
| 412 | apply (subst hoare_lemma12 [THEN mp]) | |
| 413 | apply assumption | |
| 414 | apply (rule conjI) | |
| 415 | apply (rule refl) | |
| 416 | apply assumption | |
| 417 | apply (subst hoare_lemma22) | |
| 418 | apply (subst hoare_lemma28) | |
| 419 | apply assumption | |
| 420 | apply (rule refl) | |
| 421 | apply (rule sym) | |
| 422 | apply (subst hoare_lemma27 [THEN mp]) | |
| 423 | apply assumption | |
| 424 | apply (rule conjI) | |
| 425 | apply (rule refl) | |
| 426 | apply assumption | |
| 427 | apply (rule refl) | |
| 428 | done | |
| 429 | ||
| 430 | (* ------ the main proof q o p = q ------ *) | |
| 431 | ||
| 432 | theorem hoare_main: "q oo p = q" | |
| 433 | apply (rule ext_cfun) | |
| 434 | apply (subst cfcomp2) | |
| 435 | apply (rule hoare_lemma3 [THEN disjE]) | |
| 436 | apply (erule hoare_lemma23) | |
| 437 | apply (erule hoare_lemma29) | |
| 438 | done | |
| 17291 | 439 | |
| 244 | 440 | end |