src/HOL/IMP/VC.thy
author kleing
Sun, 09 Dec 2001 15:26:13 +0100
changeset 12434 ff2efde4574d
parent 12431 07ec657249e5
child 13596 ee5f79b210c1
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1451
diff changeset
     1
(*  Title:      HOL/IMP/VC.thy
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1451
diff changeset
     3
    Author:     Tobias Nipkow
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     5
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     6
acom: annotated commands
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     7
vc:   verification-conditions
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
     8
awp:   weakest (liberal) precondition
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     9
*)
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    10
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    11
header "Verification Conditions"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    12
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    13
theory VC = Hoare:
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    14
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    15
datatype  acom = Askip
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    16
               | Aass   loc aexp
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    17
               | Asemi  acom acom
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    18
               | Aif    bexp acom acom
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    19
               | Awhile bexp assn acom
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    20
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    21
consts
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    22
  vc :: "acom => assn => assn"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    23
  awp :: "acom => assn => assn"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    24
  vcawp :: "acom => assn => assn \<times> assn"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    25
  astrip :: "acom => com"
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    26
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 4897
diff changeset
    27
primrec
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    28
  "awp Askip Q = Q"
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    29
  "awp (Aass x a) Q = (\<lambda>s. Q(s[x\<mapsto>a s]))"
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    30
  "awp (Asemi c d) Q = awp c (awp d Q)"
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    31
  "awp (Aif b c d) Q = (\<lambda>s. (b s-->awp c Q s) & (~b s-->awp d Q s))" 
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    32
  "awp (Awhile b I c) Q = I"
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    33
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 4897
diff changeset
    34
primrec
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    35
  "vc Askip Q = (\<lambda>s. True)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    36
  "vc (Aass x a) Q = (\<lambda>s. True)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    37
  "vc (Asemi c d) Q = (\<lambda>s. vc c (awp d Q) s & vc d Q s)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    38
  "vc (Aif b c d) Q = (\<lambda>s. vc c Q s & vc d Q s)" 
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    39
  "vc (Awhile b I c) Q = (\<lambda>s. (I s & ~b s --> Q s) &
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    40
                              (I s & b s --> awp c I s) & vc c I s)"
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    41
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 4897
diff changeset
    42
primrec
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1696
diff changeset
    43
  "astrip Askip = SKIP"
9241
f961c1fdff50 disambiguated := ; added Examples (factorial)
oheimb
parents: 5183
diff changeset
    44
  "astrip (Aass x a) = (x:==a)"
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1696
diff changeset
    45
  "astrip (Asemi c d) = (astrip c;astrip d)"
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    46
  "astrip (Aif b c d) = (\<IF> b \<THEN> astrip c \<ELSE> astrip d)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    47
  "astrip (Awhile b I c) = (\<WHILE> b \<DO> astrip c)"
1451
6803ecb9b198 Added vcwp
nipkow
parents: 1447
diff changeset
    48
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    49
(* simultaneous computation of vc and awp: *)
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 4897
diff changeset
    50
primrec
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    51
  "vcawp Askip Q = (\<lambda>s. True, Q)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    52
  "vcawp (Aass x a) Q = (\<lambda>s. True, \<lambda>s. Q(s[x\<mapsto>a s]))"
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    53
  "vcawp (Asemi c d) Q = (let (vcd,wpd) = vcawp d Q;
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    54
                              (vcc,wpc) = vcawp c wpd
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    55
                          in (\<lambda>s. vcc s & vcd s, wpc))"
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    56
  "vcawp (Aif b c d) Q = (let (vcd,wpd) = vcawp d Q;
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    57
                              (vcc,wpc) = vcawp c Q
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    58
                          in (\<lambda>s. vcc s & vcd s,
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    59
                              \<lambda>s.(b s --> wpc s) & (~b s --> wpd s)))"
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    60
  "vcawp (Awhile b I c) Q = (let (vcc,wpc) = vcawp c I
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    61
                             in (\<lambda>s. (I s & ~b s --> Q s) &
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 1900
diff changeset
    62
                                     (I s & b s --> wpc s) & vcc s, I))"
1451
6803ecb9b198 Added vcwp
nipkow
parents: 1447
diff changeset
    63
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    64
(*
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    65
Soundness and completeness of vc
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    66
*)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    67
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    68
declare hoare.intros [intro]
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    69
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    70
lemma l: "!s. P s --> P s" by fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    71
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    72
lemma vc_sound: "!Q. (!s. vc c Q s) --> |- {awp c Q} astrip c {Q}"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    73
apply (induct_tac "c")
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    74
    apply (simp_all (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    75
    apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    76
   apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    77
  apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    78
 (* if *)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    79
 apply (tactic "Deepen_tac 4 1")
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    80
(* while *)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    81
apply (intro allI impI) 
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    82
apply (rule conseq)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    83
  apply (rule l)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    84
 apply (rule While)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    85
 defer
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    86
 apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    87
apply (rule_tac P="awp acom fun2" in conseq)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    88
  apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    89
 apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    90
apply fast
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    91
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    92
12434
kleing
parents: 12431
diff changeset
    93
lemma awp_mono [rule_format (no_asm)]: 
kleing
parents: 12431
diff changeset
    94
  "!P Q. (!s. P s --> Q s) --> (!s. awp c P s --> awp c Q s)"
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    95
apply (induct_tac "c")
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    96
    apply (simp_all (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    97
apply (rule allI, rule allI, rule impI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    98
apply (erule allE, erule allE, erule mp)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
    99
apply (erule allE, erule allE, erule mp, assumption)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   100
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   101
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   102
12434
kleing
parents: 12431
diff changeset
   103
lemma vc_mono [rule_format (no_asm)]: 
kleing
parents: 12431
diff changeset
   104
  "!P Q. (!s. P s --> Q s) --> (!s. vc c P s --> vc c Q s)"
12431
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   105
apply (induct_tac "c")
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   106
    apply (simp_all (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   107
apply safe
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   108
apply (erule allE,erule allE,erule impE,erule_tac [2] allE,erule_tac [2] mp) 
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   109
prefer 2 apply assumption
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   110
apply (fast elim: awp_mono)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   111
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   112
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   113
lemma vc_complete: "|- {P}c{Q} ==>  
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   114
   (? ac. astrip ac = c & (!s. vc ac Q s) & (!s. P s --> awp ac Q s))"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   115
apply (erule hoare.induct)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   116
     apply (rule_tac x = "Askip" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   117
     apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   118
    apply (rule_tac x = "Aass x a" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   119
    apply (simp (no_asm))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   120
   apply clarify
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   121
   apply (rule_tac x = "Asemi ac aca" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   122
   apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   123
   apply (fast elim!: awp_mono vc_mono)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   124
  apply clarify
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   125
  apply (rule_tac x = "Aif b ac aca" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   126
  apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   127
 apply clarify
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   128
 apply (rule_tac x = "Awhile b P ac" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   129
 apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   130
apply safe
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   131
apply (rule_tac x = "ac" in exI)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   132
apply (simp (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   133
apply (fast elim!: awp_mono vc_mono)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   134
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   135
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   136
lemma vcawp_vc_awp: "!Q. vcawp c Q = (vc c Q, awp c Q)"
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   137
apply (induct_tac "c")
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   138
apply (simp_all (no_asm_simp) add: Let_def)
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   139
done
07ec657249e5 converted to Isar
kleing
parents: 9241
diff changeset
   140
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
   141
end