src/HOL/IMP/Hoare_Examples.thy
author haftmann
Tue, 19 Nov 2013 10:05:53 +0100
changeset 54489 03ff4d1e6784
parent 52554 19764bef2730
child 58410 6d46ad54a2ab
permissions -rw-r--r--
eliminiated neg_numeral in favour of - (numeral _)
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
44177
b4b5cbca2519 IMP/Util distinguishes between sets and functions again; imported only where used.
kleing
parents: 43158
diff changeset
     3
theory Hoare_Examples imports Hoare begin
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     4
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
     5
text{* Summing up the first @{text x} natural numbers in variable @{text y}. *}
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
     7
fun sum :: "int \<Rightarrow> int" where
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
     8
"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
     9
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    10
lemma sum_simps[simp]:
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    11
  "0 < i \<Longrightarrow> sum i = sum (i - 1) + i"
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    12
  "i \<le> 0 \<Longrightarrow> sum i = 0"
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    13
by(simp_all)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    14
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    15
declare sum.simps[simp del]
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    16
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    17
abbreviation "wsum ==
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    18
  WHILE Less (N 0) (V ''x'')
52554
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    19
  DO (''y'' ::= Plus (V ''y'') (V ''x'');;
19764bef2730 tuned proofs
nipkow
parents: 52529
diff changeset
    20
      ''x'' ::= Plus (V ''x'') (N -1))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    21
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    22
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    23
subsubsection{* Proof by Operational Semantics *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    24
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    25
text{* The behaviour of the loop is proved by induction: *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    26
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    27
lemma while_sum:
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    28
  "(wsum, s) \<Rightarrow> t \<Longrightarrow> t ''y'' = s ''y'' + sum(s ''x'')"
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    29
apply(induction wsum s t rule: big_step_induct)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    30
apply(auto)
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    31
done
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    32
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    33
text{* We were lucky that the proof was automatic, except for the
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    34
induction. In general, such proofs will not be so easy. The automation is
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    35
partly due to the right inversion rules that we set up as automatic
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    36
elimination rules that decompose big-step premises.
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    37
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    38
Now we prefix the loop with the necessary initialization: *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    39
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    40
lemma sum_via_bigstep:
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    41
  assumes "(''y'' ::= N 0;; wsum, s) \<Rightarrow> t"
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    42
  shows "t ''y'' = sum (s ''x'')"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    43
proof -
52108
06db08182c4b simplified example and proof
nipkow
parents: 52046
diff changeset
    44
  from assms have "(wsum,s(''y'':=0)) \<Rightarrow> t" by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    45
  from while_sum[OF this] show ?thesis by simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    46
qed
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    47
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    48
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    49
subsubsection{* Proof by Hoare Logic *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    50
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    51
text{* Note that we deal with sequences of commands from right to left,
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    52
pulling back the postcondition towards the precondition. *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    53
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    54
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
    55
apply(rule Seq)
52149
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    56
 prefer 2
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    57
 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
    58
  apply(rule Seq)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    59
   prefer 2
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    60
   apply(rule Assign)
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    61
  apply(rule Assign')
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    62
  apply simp
32b1dbda331c simpler proof through custom summation function
nipkow
parents: 52108
diff changeset
    63
 apply(simp)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    64
apply(rule Assign')
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    65
apply simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    66
done
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    67
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    68
text{* The proof is intentionally an apply skript because it merely composes
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    69
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
    70
be proved. But since those proofs are 1-liners, a structured proof is
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    71
overkill. In fact, we shall learn later that the application of the Hoare
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    72
rules can be automated completely and all that is left for the user is to
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    73
provide the loop invariants and prove the side-conditions. *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    74
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    75
end