src/HOL/IMP/Hoare.thy
author blanchet
Mon, 14 Dec 2009 12:14:12 +0100
changeset 34120 f9920a3ddf50
parent 27362 a6dc1769fdda
child 35735 f139a9bb6501
permissions -rw-r--r--
added "no_assms" option to Refute, and include structured proof assumptions by default; will do the same for Quickcheck unless there are objections
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1447
diff changeset
     1
(*  Title:      HOL/IMP/Hoare.thy
938
621be7ec81d7 *** empty log message ***
nipkow
parents: 937
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1447
diff changeset
     3
    Author:     Tobias Nipkow
936
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     4
    Copyright   1995 TUM
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     5
*)
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     6
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
     7
header "Inductive Definition of Hoare Logic"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13630
diff changeset
     9
theory Hoare imports Denotation begin
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents: 1374
diff changeset
    10
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    11
types assn = "state => bool"
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents: 1374
diff changeset
    12
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    13
definition
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    14
  hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    15
  "|= {P}c{Q} = (!s t. (s,t) : C(c) --> P s --> Q t)"
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    16
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    17
inductive
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    18
  hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    19
where
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    20
  skip: "|- {P}\<SKIP>{P}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    21
| ass:  "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    22
| semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    23
| If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==>
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    24
      |- {P} \<IF> b \<THEN> c \<ELSE> d {Q}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    25
| While: "|- {%s. P s & b s} c {P} ==>
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    26
         |- {P} \<WHILE> b \<DO> c {%s. P s & ~b s}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    27
| conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==>
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1481
diff changeset
    28
          |- {P'}c{Q'}"
1481
03f096efa26d Modified datatype com.
nipkow
parents: 1476
diff changeset
    29
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    30
definition
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    31
  wp :: "com => assn => assn" where
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    32
  "wp c Q = (%s. !t. (s,t) : C(c) --> Q t)"
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    33
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    34
(*
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    35
Soundness (and part of) relative completeness of Hoare rules
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    36
wrt denotational semantics
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    37
*)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    38
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    39
lemma hoare_conseq1: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    40
apply (erule hoare.conseq)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    41
apply  assumption
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    42
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    43
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    44
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    45
lemma hoare_conseq2: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    46
apply (rule hoare.conseq)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    47
prefer 2 apply    (assumption)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    48
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    49
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    50
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    51
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    52
lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    53
apply (unfold hoare_valid_def)
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    54
apply (induct set: hoare)
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    55
     apply (simp_all (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    56
  apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    57
 apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    58
apply (rule allI, rule allI, rule impI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    59
apply (erule lfp_induct2)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    60
 apply (rule Gamma_mono)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    61
apply (unfold Gamma_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    62
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    63
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    64
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    65
lemma wp_SKIP: "wp \<SKIP> Q = Q"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    66
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    67
apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    68
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    69
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    70
lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\<mapsto>a s]))"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    71
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    72
apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    73
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    74
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    75
lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    76
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    77
apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    78
apply (rule ext)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    79
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    80
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    81
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    82
lemma wp_If:
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    83
 "wp (\<IF> b \<THEN> c \<ELSE> d) Q = (%s. (b s --> wp c Q s) &  (~b s --> wp d Q s))"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    84
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    85
apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    86
apply (rule ext)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    87
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    88
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    89
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    90
lemma wp_While_True:
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    91
  "b s ==> wp (\<WHILE> b \<DO> c) Q s = wp (c;\<WHILE> b \<DO> c) Q s"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    92
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    93
apply (subst C_While_If)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    94
apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    95
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    96
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    97
lemma wp_While_False: "~b s ==> wp (\<WHILE> b \<DO> c) Q s = Q s"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    98
apply (unfold wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    99
apply (subst C_While_If)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   100
apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   101
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   102
12434
kleing
parents: 12431
diff changeset
   103
lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   104
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   105
(*Not suitable for rewriting: LOOPS!*)
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   106
lemma wp_While_if:
12434
kleing
parents: 12431
diff changeset
   107
  "wp (\<WHILE> b \<DO> c) Q s = (if b s then wp (c;\<WHILE> b \<DO> c) Q s else Q s)"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   108
  by simp
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   109
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   110
lemma wp_While: "wp (\<WHILE> b \<DO> c) Q s =
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   111
   (s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   112
apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   113
apply (rule iffI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   114
 apply (rule weak_coinduct)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   115
  apply (erule CollectI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   116
 apply safe
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   117
  apply simp
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   118
 apply simp
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   119
apply (simp add: wp_def Gamma_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   120
apply (intro strip)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   121
apply (rule mp)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   122
 prefer 2 apply (assumption)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   123
apply (erule lfp_induct2)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   124
apply (fast intro!: monoI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   125
apply (subst gfp_unfold)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   126
 apply (fast intro!: monoI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   127
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   128
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   129
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   130
declare C_while [simp del]
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   131
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   132
lemmas [intro!] = hoare.skip hoare.ass hoare.semi hoare.If
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   133
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   134
lemma wp_is_pre: "|- {wp c Q} c {Q}"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18372
diff changeset
   135
apply (induct c arbitrary: Q)
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   136
    apply (simp_all (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   137
    apply fast+
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   138
 apply (blast intro: hoare_conseq1)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   139
apply (rule hoare_conseq2)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   140
 apply (rule hoare.While)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   141
 apply (rule hoare_conseq1)
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   142
  prefer 2 apply fast
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   143
  apply safe
13630
a013a9dd370f Got rid of rotates because of new simplifier
nipkow
parents: 12546
diff changeset
   144
 apply simp
a013a9dd370f Got rid of rotates because of new simplifier
nipkow
parents: 12546
diff changeset
   145
apply simp
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   146
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   147
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   148
lemma hoare_relative_complete: "|= {P}c{Q} ==> |- {P}c{Q}"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   149
apply (rule hoare_conseq1 [OF _ wp_is_pre])
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   150
apply (unfold hoare_valid_def wp_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   151
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   152
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   153
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
   154
end