| 1476 |      1 | (*  Title:      HOL/IMP/VC.thy
 | 
| 1447 |      2 |     ID:         $Id$
 | 
| 1476 |      3 |     Author:     Tobias Nipkow
 | 
| 1447 |      4 |     Copyright   1996 TUM
 | 
|  |      5 | 
 | 
|  |      6 | acom: annotated commands
 | 
|  |      7 | vc:   verification-conditions
 | 
| 2810 |      8 | awp:   weakest (liberal) precondition
 | 
| 1447 |      9 | *)
 | 
|  |     10 | 
 | 
|  |     11 | VC  =  Hoare +
 | 
|  |     12 | 
 | 
|  |     13 | datatype  acom = Askip
 | 
|  |     14 |                | Aass   loc aexp
 | 
|  |     15 |                | Asemi  acom acom
 | 
|  |     16 |                | Aif    bexp acom acom
 | 
|  |     17 |                | Awhile bexp assn acom
 | 
|  |     18 | 
 | 
|  |     19 | consts
 | 
| 2810 |     20 |   vc,awp :: acom => assn => assn
 | 
|  |     21 |   vcawp :: "acom => assn => assn * assn"
 | 
| 1447 |     22 |   astrip :: acom => com
 | 
|  |     23 | 
 | 
| 2810 |     24 | primrec awp acom
 | 
|  |     25 |   "awp Askip Q = Q"
 | 
|  |     26 |   "awp (Aass x a) Q = (%s.Q(s[a s/x]))"
 | 
|  |     27 |   "awp (Asemi c d) Q = awp c (awp d Q)"
 | 
|  |     28 |   "awp (Aif b c d) Q = (%s. (b s-->awp c Q s) & (~b s-->awp d Q s))" 
 | 
|  |     29 |   "awp (Awhile b I c) Q = I"
 | 
| 1447 |     30 | 
 | 
|  |     31 | primrec vc acom
 | 
| 1900 |     32 |   "vc Askip Q = (%s.True)"
 | 
|  |     33 |   "vc (Aass x a) Q = (%s.True)"
 | 
| 2810 |     34 |   "vc (Asemi c d) Q = (%s. vc c (awp d Q) s & vc d Q s)"
 | 
| 1900 |     35 |   "vc (Aif b c d) Q = (%s. vc c Q s & vc d Q s)" 
 | 
|  |     36 |   "vc (Awhile b I c) Q = (%s. (I s & ~b s --> Q s) &
 | 
| 2810 |     37 |                               (I s & b s --> awp c I s) & vc c I s)"
 | 
| 1447 |     38 | 
 | 
|  |     39 | primrec astrip acom
 | 
| 1900 |     40 |   "astrip Askip = SKIP"
 | 
|  |     41 |   "astrip (Aass x a) = (x:=a)"
 | 
|  |     42 |   "astrip (Asemi c d) = (astrip c;astrip d)"
 | 
|  |     43 |   "astrip (Aif b c d) = (IF b THEN astrip c ELSE astrip d)"
 | 
|  |     44 |   "astrip (Awhile b I c) = (WHILE b DO astrip c)"
 | 
| 1451 |     45 | 
 | 
| 2810 |     46 | (* simultaneous computation of vc and awp: *)
 | 
|  |     47 | primrec vcawp acom
 | 
|  |     48 |   "vcawp Askip Q = (%s.True, Q)"
 | 
|  |     49 |   "vcawp (Aass x a) Q = (%s.True, %s.Q(s[a s/x]))"
 | 
|  |     50 |   "vcawp (Asemi c d) Q = (let (vcd,wpd) = vcawp d Q;
 | 
|  |     51 |                               (vcc,wpc) = vcawp c wpd
 | 
|  |     52 |                           in (%s. vcc s & vcd s, wpc))"
 | 
|  |     53 |   "vcawp (Aif b c d) Q = (let (vcd,wpd) = vcawp d Q;
 | 
|  |     54 |                               (vcc,wpc) = vcawp c Q
 | 
|  |     55 |                           in (%s. vcc s & vcd s,
 | 
|  |     56 |                               %s.(b s --> wpc s) & (~b s --> wpd s)))"
 | 
|  |     57 |   "vcawp (Awhile b I c) Q = (let (vcc,wpc) = vcawp c I
 | 
|  |     58 |                              in (%s. (I s & ~b s --> Q s) &
 | 
|  |     59 |                                      (I s & b s --> wpc s) & vcc s, I))"
 | 
| 1451 |     60 | 
 | 
| 1447 |     61 | end
 |