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