| 1476 |      1 | (*  Title:      HOL/IMP/Hoare.thy
 | 
| 938 |      2 |     ID:         $Id$
 | 
| 1476 |      3 |     Author:     Tobias Nipkow
 | 
| 936 |      4 |     Copyright   1995 TUM
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 12431 |      7 | header "Inductive Definition of Hoare Logic"
 | 
|  |      8 | 
 | 
| 16417 |      9 | theory Hoare imports Denotation begin
 | 
| 1447 |     10 | 
 | 
| 12431 |     11 | types assn = "state => bool"
 | 
| 1447 |     12 | 
 | 
| 12431 |     13 | constdefs hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50)
 | 
| 1696 |     14 |           "|= {P}c{Q} == !s t. (s,t) : C(c) --> P s --> Q t"
 | 
| 939 |     15 | 
 | 
| 1696 |     16 | consts hoare :: "(assn * com * assn) set"
 | 
| 12546 |     17 | syntax "_hoare" :: "[bool,com,bool] => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
 | 
| 1486 |     18 | translations "|- {P}c{Q}" == "(P,c,Q) : hoare"
 | 
| 939 |     19 | 
 | 
| 1789 |     20 | inductive hoare
 | 
| 12431 |     21 | intros
 | 
|  |     22 |   skip: "|- {P}\<SKIP>{P}"
 | 
|  |     23 |   ass:  "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
 | 
|  |     24 |   semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
 | 
|  |     25 |   If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==>
 | 
|  |     26 |       |- {P} \<IF> b \<THEN> c \<ELSE> d {Q}"
 | 
|  |     27 |   While: "|- {%s. P s & b s} c {P} ==>
 | 
|  |     28 |          |- {P} \<WHILE> b \<DO> c {%s. P s & ~b s}"
 | 
|  |     29 |   conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==>
 | 
| 1486 |     30 |           |- {P'}c{Q'}"
 | 
| 1481 |     31 | 
 | 
| 12431 |     32 | constdefs wp :: "com => assn => assn"
 | 
| 2810 |     33 |           "wp c Q == (%s. !t. (s,t) : C(c) --> Q t)"
 | 
| 939 |     34 | 
 | 
| 18372 |     35 | (*
 | 
| 12431 |     36 | Soundness (and part of) relative completeness of Hoare rules
 | 
|  |     37 | wrt denotational semantics
 | 
|  |     38 | *)
 | 
|  |     39 | 
 | 
|  |     40 | lemma hoare_conseq1: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
 | 
|  |     41 | apply (erule hoare.conseq)
 | 
|  |     42 | apply  assumption
 | 
|  |     43 | apply fast
 | 
|  |     44 | done
 | 
|  |     45 | 
 | 
|  |     46 | lemma hoare_conseq2: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
 | 
|  |     47 | apply (rule hoare.conseq)
 | 
|  |     48 | prefer 2 apply    (assumption)
 | 
|  |     49 | apply fast
 | 
|  |     50 | apply fast
 | 
|  |     51 | done
 | 
|  |     52 | 
 | 
|  |     53 | lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}"
 | 
|  |     54 | apply (unfold hoare_valid_def)
 | 
| 18372 |     55 | apply (induct set: hoare)
 | 
| 12431 |     56 |      apply (simp_all (no_asm_simp))
 | 
|  |     57 |   apply fast
 | 
|  |     58 |  apply fast
 | 
|  |     59 | apply (rule allI, rule allI, rule impI)
 | 
|  |     60 | apply (erule lfp_induct2)
 | 
|  |     61 |  apply (rule Gamma_mono)
 | 
|  |     62 | apply (unfold Gamma_def)
 | 
|  |     63 | apply fast
 | 
|  |     64 | done
 | 
|  |     65 | 
 | 
|  |     66 | lemma wp_SKIP: "wp \<SKIP> Q = Q"
 | 
|  |     67 | apply (unfold wp_def)
 | 
|  |     68 | apply (simp (no_asm))
 | 
|  |     69 | done
 | 
|  |     70 | 
 | 
|  |     71 | lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\<mapsto>a s]))"
 | 
|  |     72 | apply (unfold wp_def)
 | 
|  |     73 | apply (simp (no_asm))
 | 
|  |     74 | done
 | 
|  |     75 | 
 | 
|  |     76 | lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)"
 | 
|  |     77 | apply (unfold wp_def)
 | 
|  |     78 | apply (simp (no_asm))
 | 
|  |     79 | apply (rule ext)
 | 
|  |     80 | apply fast
 | 
|  |     81 | done
 | 
|  |     82 | 
 | 
| 18372 |     83 | lemma wp_If:
 | 
| 12431 |     84 |  "wp (\<IF> b \<THEN> c \<ELSE> d) Q = (%s. (b s --> wp c Q s) &  (~b s --> wp d Q s))"
 | 
|  |     85 | apply (unfold wp_def)
 | 
|  |     86 | apply (simp (no_asm))
 | 
|  |     87 | apply (rule ext)
 | 
|  |     88 | apply fast
 | 
|  |     89 | done
 | 
|  |     90 | 
 | 
| 18372 |     91 | lemma wp_While_True:
 | 
| 12431 |     92 |   "b s ==> wp (\<WHILE> b \<DO> c) Q s = wp (c;\<WHILE> b \<DO> c) Q s"
 | 
|  |     93 | apply (unfold wp_def)
 | 
|  |     94 | apply (subst C_While_If)
 | 
|  |     95 | apply (simp (no_asm_simp))
 | 
|  |     96 | done
 | 
|  |     97 | 
 | 
|  |     98 | lemma wp_While_False: "~b s ==> wp (\<WHILE> b \<DO> c) Q s = Q s"
 | 
|  |     99 | apply (unfold wp_def)
 | 
|  |    100 | apply (subst C_While_If)
 | 
|  |    101 | apply (simp (no_asm_simp))
 | 
|  |    102 | done
 | 
|  |    103 | 
 | 
| 12434 |    104 | lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False
 | 
| 12431 |    105 | 
 | 
|  |    106 | (*Not suitable for rewriting: LOOPS!*)
 | 
| 18372 |    107 | lemma wp_While_if:
 | 
| 12434 |    108 |   "wp (\<WHILE> b \<DO> c) Q s = (if b s then wp (c;\<WHILE> b \<DO> c) Q s else Q s)"
 | 
| 18372 |    109 |   by simp
 | 
| 12431 |    110 | 
 | 
| 18372 |    111 | lemma wp_While: "wp (\<WHILE> b \<DO> c) Q s =
 | 
| 12431 |    112 |    (s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))"
 | 
|  |    113 | apply (simp (no_asm))
 | 
|  |    114 | apply (rule iffI)
 | 
|  |    115 |  apply (rule weak_coinduct)
 | 
|  |    116 |   apply (erule CollectI)
 | 
|  |    117 |  apply safe
 | 
|  |    118 |   apply simp
 | 
|  |    119 |  apply simp
 | 
|  |    120 | apply (simp add: wp_def Gamma_def)
 | 
|  |    121 | apply (intro strip)
 | 
|  |    122 | apply (rule mp)
 | 
|  |    123 |  prefer 2 apply (assumption)
 | 
|  |    124 | apply (erule lfp_induct2)
 | 
|  |    125 | apply (fast intro!: monoI)
 | 
|  |    126 | apply (subst gfp_unfold)
 | 
|  |    127 |  apply (fast intro!: monoI)
 | 
|  |    128 | apply fast
 | 
|  |    129 | done
 | 
|  |    130 | 
 | 
|  |    131 | declare C_while [simp del]
 | 
|  |    132 | 
 | 
| 18372 |    133 | lemmas [intro!] = hoare.skip hoare.ass hoare.semi hoare.If
 | 
| 12431 |    134 | 
 | 
| 18372 |    135 | lemma wp_is_pre: "|- {wp c Q} c {Q}"
 | 
| 20503 |    136 | apply (induct c arbitrary: Q)
 | 
| 12431 |    137 |     apply (simp_all (no_asm))
 | 
|  |    138 |     apply fast+
 | 
|  |    139 |  apply (blast intro: hoare_conseq1)
 | 
|  |    140 | apply (rule hoare_conseq2)
 | 
|  |    141 |  apply (rule hoare.While)
 | 
|  |    142 |  apply (rule hoare_conseq1)
 | 
| 18372 |    143 |   prefer 2 apply fast
 | 
| 12431 |    144 |   apply safe
 | 
| 13630 |    145 |  apply simp
 | 
|  |    146 | apply simp
 | 
| 12431 |    147 | done
 | 
|  |    148 | 
 | 
|  |    149 | lemma hoare_relative_complete: "|= {P}c{Q} ==> |- {P}c{Q}"
 | 
|  |    150 | apply (rule hoare_conseq1 [OF _ wp_is_pre])
 | 
|  |    151 | apply (unfold hoare_valid_def wp_def)
 | 
|  |    152 | apply fast
 | 
|  |    153 | done
 | 
|  |    154 | 
 | 
| 939 |    155 | end
 |