src/HOL/IMP/Hoare_Examples.thy
author haftmann
Fri, 24 Aug 2018 20:22:14 +0000
changeset 68802 3974935e0252
parent 68776 403dd13cf6e9
child 69505 cc2d676d5395
permissions -rw-r--r--
some modernization of notation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     2
68776
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     3
subsection "Examples"
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     4
44177
b4b5cbca2519 IMP/Util distinguishes between sets and functions again; imported only where used.
kleing
parents: 43158
diff changeset
     5
theory Hoare_Examples imports Hoare begin
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
64851
33aab75ff423 hide global sum
nipkow
parents: 61028
diff changeset
     7
hide_const (open) sum
33aab75ff423 hide global sum
nipkow
parents: 61028
diff changeset
     8
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
     9
text\<open>Summing up the first @{text x} natural numbers in variable @{text y}.\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    10
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    11
fun sum :: "int \<Rightarrow> int" where
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    12
"sum i = (if i \<le> 0 then 0 else sum (i - 1) + i)"
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    13
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    14
lemma sum_simps[simp]:
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    15
  "0 < i \<Longrightarrow> sum i = sum (i - 1) + i"
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    16
  "i \<le> 0 \<Longrightarrow> sum i = 0"
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    17
by(simp_all)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    18
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    19
declare sum.simps[simp del]
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    20
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    21
abbreviation "wsum ==
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    22
  WHILE Less (N 0) (V ''x'')
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    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: 54489
diff changeset
    24
      ''x'' ::= Plus (V ''x'') (N (- 1)))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    25
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    26
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    27
subsubsection\<open>Proof by Operational Semantics\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    28
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    29
text\<open>The behaviour of the loop is proved by induction:\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    31
lemma while_sum:
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    32
  "(wsum, s) \<Rightarrow> t \<Longrightarrow> t ''y'' = s ''y'' + sum(s ''x'')"
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    33
apply(induction wsum s t rule: big_step_induct)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    34
apply(auto)
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    35
done
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    36
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    37
text\<open>We were lucky that the proof was automatic, except for the
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    38
induction. In general, such proofs will not be so easy. The automation is
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    39
partly due to the right inversion rules that we set up as automatic
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    40
elimination rules that decompose big-step premises.
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    41
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    42
Now we prefix the loop with the necessary initialization:\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    43
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    44
lemma sum_via_bigstep:
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    45
  assumes "(''y'' ::= N 0;; wsum, s) \<Rightarrow> t"
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    46
  shows "t ''y'' = sum (s ''x'')"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    47
proof -
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    48
  from assms have "(wsum,s(''y'':=0)) \<Rightarrow> t" by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    49
  from while_sum[OF this] show ?thesis by simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    50
qed
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    51
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    52
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    53
subsubsection\<open>Proof by Hoare Logic\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    54
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    55
text\<open>Note that we deal with sequences of commands from right to left,
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    56
pulling back the postcondition towards the precondition.\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    57
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    58
lemma "\<turnstile> {\<lambda>s. s ''x'' = n} ''y'' ::= N 0;; wsum {\<lambda>s. s ''y'' = sum n}"
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45015
diff changeset
    59
apply(rule Seq)
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    60
 prefer 2
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    61
 apply(rule While' [where P = "\<lambda>s. (s ''y'' = sum n - sum(s ''x''))"])
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    62
  apply(rule Seq)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    63
   prefer 2
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    64
   apply(rule Assign)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    65
  apply(rule Assign')
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    66
  apply simp
61028
99d58362eeeb fixed typo in comment
blanchet
parents: 58410
diff changeset
    67
 apply simp
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    68
apply(rule Assign')
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    69
apply simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    70
done
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    71
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    72
text\<open>The proof is intentionally an apply script because it merely composes
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    73
the rules of Hoare logic. Of course, in a few places side conditions have to
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    74
be proved. But since those proofs are 1-liners, a structured proof is
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    75
overkill. In fact, we shall learn later that the application of the Hoare
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    76
rules can be automated completely and all that is left for the user is to
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64851
diff changeset
    77
provide the loop invariants and prove the side-conditions.\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    78
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    79
end