| author | haftmann | 
| Sun, 09 Feb 2020 21:58:42 +0000 | |
| changeset 71425 | f2da99316b86 | 
| parent 69505 | cc2d676d5395 | 
| permissions | -rw-r--r-- | 
| 43158 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 68776 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 3 | subsection "Examples" | 
| 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 4 | |
| 44177 
b4b5cbca2519
IMP/Util distinguishes between sets and functions again; imported only where used.
 kleing parents: 
43158diff
changeset | 5 | theory Hoare_Examples imports Hoare begin | 
| 43158 | 6 | |
| 64851 | 7 | hide_const (open) sum | 
| 8 | ||
| 69505 | 9 | text\<open>Summing up the first \<open>x\<close> natural numbers in variable \<open>y\<close>.\<close> | 
| 43158 | 10 | |
| 52149 | 11 | fun sum :: "int \<Rightarrow> int" where | 
| 52554 | 12 | "sum i = (if i \<le> 0 then 0 else sum (i - 1) + i)" | 
| 52149 | 13 | |
| 52554 | 14 | lemma sum_simps[simp]: | 
| 15 | "0 < i \<Longrightarrow> sum i = sum (i - 1) + i" | |
| 16 | "i \<le> 0 \<Longrightarrow> sum i = 0" | |
| 52149 | 17 | by(simp_all) | 
| 18 | ||
| 19 | declare sum.simps[simp del] | |
| 20 | ||
| 52108 | 21 | abbreviation "wsum == | 
| 22 | WHILE Less (N 0) (V ''x'') | |
| 52554 | 23 |   DO (''y'' ::= Plus (V ''y'') (V ''x'');;
 | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
54489diff
changeset | 24 | ''x'' ::= Plus (V ''x'') (N (- 1)))" | 
| 43158 | 25 | |
| 26 | ||
| 67406 | 27 | subsubsection\<open>Proof by Operational Semantics\<close> | 
| 43158 | 28 | |
| 67406 | 29 | text\<open>The behaviour of the loop is proved by induction:\<close> | 
| 43158 | 30 | |
| 52149 | 31 | lemma while_sum: | 
| 32 | "(wsum, s) \<Rightarrow> t \<Longrightarrow> t ''y'' = s ''y'' + sum(s ''x'')" | |
| 33 | apply(induction wsum s t rule: big_step_induct) | |
| 34 | apply(auto) | |
| 52108 | 35 | done | 
| 36 | ||
| 67406 | 37 | text\<open>We were lucky that the proof was automatic, except for the | 
| 43158 | 38 | induction. In general, such proofs will not be so easy. The automation is | 
| 39 | partly due to the right inversion rules that we set up as automatic | |
| 40 | elimination rules that decompose big-step premises. | |
| 41 | ||
| 67406 | 42 | Now we prefix the loop with the necessary initialization:\<close> | 
| 43158 | 43 | |
| 44 | lemma sum_via_bigstep: | |
| 52149 | 45 |   assumes "(''y'' ::= N 0;; wsum, s) \<Rightarrow> t"
 | 
| 46 | shows "t ''y'' = sum (s ''x'')" | |
| 43158 | 47 | proof - | 
| 52108 | 48 |   from assms have "(wsum,s(''y'':=0)) \<Rightarrow> t" by auto
 | 
| 43158 | 49 | from while_sum[OF this] show ?thesis by simp | 
| 50 | qed | |
| 51 | ||
| 52 | ||
| 67406 | 53 | subsubsection\<open>Proof by Hoare Logic\<close> | 
| 43158 | 54 | |
| 67406 | 55 | text\<open>Note that we deal with sequences of commands from right to left, | 
| 56 | pulling back the postcondition towards the precondition.\<close> | |
| 43158 | 57 | |
| 52149 | 58 | lemma "\<turnstile> {\<lambda>s. s ''x'' = n} ''y'' ::= N 0;; wsum {\<lambda>s. s ''y'' = sum n}"
 | 
| 47818 | 59 | apply(rule Seq) | 
| 52149 | 60 | prefer 2 | 
| 61 | apply(rule While' [where P = "\<lambda>s. (s ''y'' = sum n - sum(s ''x''))"]) | |
| 62 | apply(rule Seq) | |
| 63 | prefer 2 | |
| 64 | apply(rule Assign) | |
| 65 | apply(rule Assign') | |
| 66 | apply simp | |
| 61028 | 67 | apply simp | 
| 43158 | 68 | apply(rule Assign') | 
| 69 | apply simp | |
| 70 | done | |
| 71 | ||
| 67406 | 72 | text\<open>The proof is intentionally an apply script because it merely composes | 
| 43158 | 73 | the rules of Hoare logic. Of course, in a few places side conditions have to | 
| 74 | be proved. But since those proofs are 1-liners, a structured proof is | |
| 75 | overkill. In fact, we shall learn later that the application of the Hoare | |
| 76 | rules can be automated completely and all that is left for the user is to | |
| 67406 | 77 | provide the loop invariants and prove the side-conditions.\<close> | 
| 43158 | 78 | |
| 79 | end |