src/HOL/IMP/Hoare.thy
author wenzelm
Sun, 28 Nov 2010 14:01:20 +0100
changeset 40781 ba5be5c3d477
parent 35754 8e7dba5f00f5
child 41589 bbd861837ebc
permissions -rw-r--r--
updated versions;
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
35754
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
     9
theory Hoare imports Natural 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
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    13
inductive
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    14
  hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    15
where
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    16
  skip: "|- {P}\<SKIP>{P}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    17
| ass:  "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    18
| semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    19
| 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
    20
      |- {P} \<IF> b \<THEN> c \<ELSE> d {Q}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    21
| While: "|- {%s. P s & b s} c {P} ==>
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    22
         |- {P} \<WHILE> b \<DO> c {%s. P s & ~b s}"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 20503
diff changeset
    23
| 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
    24
          |- {P'}c{Q'}"
1481
03f096efa26d Modified datatype com.
nipkow
parents: 1476
diff changeset
    25
35735
f139a9bb6501 converted proofs to Isar
nipkow
parents: 27362
diff changeset
    26
lemma strengthen_pre: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
f139a9bb6501 converted proofs to Isar
nipkow
parents: 27362
diff changeset
    27
by (blast intro: conseq)
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    28
35735
f139a9bb6501 converted proofs to Isar
nipkow
parents: 27362
diff changeset
    29
lemma weaken_post: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
f139a9bb6501 converted proofs to Isar
nipkow
parents: 27362
diff changeset
    30
by (blast intro: conseq)
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    31
35754
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
    32
lemma While':
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
    33
assumes "|- {%s. P s & b s} c {P}" and "ALL s. P s & \<not> b s \<longrightarrow> Q s"
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
    34
shows "|- {P} \<WHILE> b \<DO> c {Q}"
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
    35
by(rule weaken_post[OF While[OF assms(1)] assms(2)])
35735
f139a9bb6501 converted proofs to Isar
nipkow
parents: 27362
diff changeset
    36
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    37
35754
8e7dba5f00f5 Reorganized Hoare logic theories; added Hoare_Den
nipkow
parents: 35735
diff changeset
    38
lemmas [simp] = skip ass semi If
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    39
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    40
lemmas [intro!] = hoare.skip hoare.ass hoare.semi hoare.If
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    41
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    42
end