| 8177 |      1 | (*  Title:      HOL/IMPP/Hoare.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     David von Oheimb
 | 
|  |      4 |     Copyright   1999 TUM
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 17477 |      7 | header {* Inductive definition of Hoare logic for partial correctness *}
 | 
|  |      8 | 
 | 
|  |      9 | theory Hoare
 | 
|  |     10 | imports Natural
 | 
|  |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | text {*
 | 
|  |     14 |   Completeness is taken relative to completeness of the underlying logic.
 | 
|  |     15 | 
 | 
|  |     16 |   Two versions of completeness proof: nested single recursion
 | 
|  |     17 |   vs. simultaneous recursion in call rule
 | 
|  |     18 | *}
 | 
| 8177 |     19 | 
 | 
|  |     20 | types 'a assn = "'a => state => bool"
 | 
|  |     21 | translations
 | 
| 17477 |     22 |   "a assn"   <= (type)"a => state => bool"
 | 
| 8177 |     23 | 
 | 
|  |     24 | constdefs
 | 
|  |     25 |   state_not_singleton :: bool
 | 
| 17477 |     26 |   "state_not_singleton == \<exists>s t::state. s ~= t" (* at least two elements *)
 | 
| 8177 |     27 | 
 | 
|  |     28 |   peek_and    :: "'a assn => (state => bool) => 'a assn" (infixr "&>" 35)
 | 
| 17477 |     29 |   "peek_and P p == %Z s. P Z s & p s"
 | 
| 8177 |     30 | 
 | 
|  |     31 | datatype 'a triple =
 | 
| 17477 |     32 |   triple "'a assn"  com  "'a assn"       ("{(1_)}./ (_)/ .{(1_)}" [3,60,3] 58)
 | 
|  |     33 | 
 | 
| 8177 |     34 | consts
 | 
| 17477 |     35 |   triple_valid ::            "nat => 'a triple     => bool" ( "|=_:_" [0 , 58] 57)
 | 
|  |     36 |   hoare_valids ::  "'a triple set => 'a triple set => bool" ("_||=_"  [58, 58] 57)
 | 
|  |     37 |   hoare_derivs :: "('a triple set *  'a triple set)   set"
 | 
| 8177 |     38 | syntax
 | 
| 17477 |     39 |   triples_valid::            "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57)
 | 
|  |     40 |   hoare_valid  ::  "'a triple set => 'a triple     => bool" ("_|=_"   [58, 58] 57)
 | 
|  |     41 | "@hoare_derivs"::  "'a triple set => 'a triple set => bool" ("_||-_"  [58, 58] 57)
 | 
|  |     42 | "@hoare_deriv" ::  "'a triple set => 'a triple     => bool" ("_|-_"   [58, 58] 57)
 | 
| 8177 |     43 | 
 | 
| 17477 |     44 | defs triple_valid_def: "|=n:t  ==  case t of {P}.c.{Q} =>
 | 
|  |     45 |                                 !Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s')"
 | 
| 8177 |     46 | translations          "||=n:G" == "Ball G (triple_valid n)"
 | 
| 17477 |     47 | defs hoare_valids_def: "G||=ts   ==  !n. ||=n:G --> ||=n:ts"
 | 
| 8177 |     48 | translations         "G |=t  " == " G||={t}"
 | 
|  |     49 |                      "G||-ts"  == "(G,ts) : hoare_derivs"
 | 
|  |     50 |                      "G |-t"   == " G||-{t}"
 | 
|  |     51 | 
 | 
|  |     52 | (* Most General Triples *)
 | 
| 17477 |     53 | constdefs MGT    :: "com => state triple"            ("{=}._.{->}" [60] 58)
 | 
| 8177 |     54 |          "{=}.c.{->} == {%Z s0. Z = s0}. c .{%Z s1. <c,Z> -c-> s1}"
 | 
|  |     55 | 
 | 
| 17477 |     56 | inductive hoare_derivs intros
 | 
|  |     57 | 
 | 
|  |     58 |   empty:    "G||-{}"
 | 
|  |     59 |   insert: "[| G |-t;  G||-ts |]
 | 
|  |     60 |         ==> G||-insert t ts"
 | 
| 8177 |     61 | 
 | 
| 17477 |     62 |   asm:      "ts <= G ==>
 | 
|  |     63 |              G||-ts" (* {P}.BODY pn.{Q} instead of (general) t for SkipD_lemma *)
 | 
| 8177 |     64 | 
 | 
| 17477 |     65 |   cut:   "[| G'||-ts; G||-G' |] ==> G||-ts" (* for convenience and efficiency *)
 | 
| 8177 |     66 | 
 | 
| 17477 |     67 |   weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts"
 | 
| 8177 |     68 | 
 | 
| 17477 |     69 |   conseq: "!Z s. P  Z  s --> (? P' Q'. G|-{P'}.c.{Q'} &
 | 
|  |     70 |                                    (!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s'))
 | 
|  |     71 |           ==> G|-{P}.c.{Q}"
 | 
| 8177 |     72 | 
 | 
|  |     73 | 
 | 
| 17477 |     74 |   Skip:  "G|-{P}. SKIP .{P}"
 | 
| 8177 |     75 | 
 | 
| 17477 |     76 |   Ass:   "G|-{%Z s. P Z (s[X::=a s])}. X:==a .{P}"
 | 
| 8177 |     77 | 
 | 
| 17477 |     78 |   Local: "G|-{P}. c .{%Z s. Q Z (s[Loc X::=s'<X>])}
 | 
|  |     79 |       ==> G|-{%Z s. s'=s & P Z (s[Loc X::=a s])}. LOCAL X:=a IN c .{Q}"
 | 
| 8177 |     80 | 
 | 
| 17477 |     81 |   Comp:  "[| G|-{P}.c.{Q};
 | 
|  |     82 |              G|-{Q}.d.{R} |]
 | 
|  |     83 |          ==> G|-{P}. (c;;d) .{R}"
 | 
| 8177 |     84 | 
 | 
| 17477 |     85 |   If:    "[| G|-{P &>        b }.c.{Q};
 | 
|  |     86 |              G|-{P &> (Not o b)}.d.{Q} |]
 | 
|  |     87 |          ==> G|-{P}. IF b THEN c ELSE d .{Q}"
 | 
| 8177 |     88 | 
 | 
| 17477 |     89 |   Loop:  "G|-{P &> b}.c.{P} ==>
 | 
|  |     90 |           G|-{P}. WHILE b DO c .{P &> (Not o b)}"
 | 
| 8177 |     91 | 
 | 
|  |     92 | (*
 | 
| 17477 |     93 |   BodyN: "(insert ({P}. BODY pn  .{Q}) G)
 | 
|  |     94 |            |-{P}.  the (body pn) .{Q} ==>
 | 
|  |     95 |           G|-{P}.       BODY pn  .{Q}"
 | 
| 8177 |     96 | *)
 | 
| 17477 |     97 |   Body:  "[| G Un (%p. {P p}.      BODY p  .{Q p})`Procs
 | 
|  |     98 |                ||-(%p. {P p}. the (body p) .{Q p})`Procs |]
 | 
|  |     99 |          ==>  G||-(%p. {P p}.      BODY p  .{Q p})`Procs"
 | 
| 8177 |    100 | 
 | 
| 17477 |    101 |   Call:     "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])}
 | 
|  |    102 |          ==> G|-{%Z s. s'=s & P Z (setlocs s newlocs[Loc Arg::=a s])}.
 | 
|  |    103 |              X:=CALL pn(a) .{Q}"
 | 
|  |    104 | 
 | 
| 19803 |    105 | 
 | 
|  |    106 | section {* Soundness and relative completeness of Hoare rules wrt operational semantics *}
 | 
|  |    107 | 
 | 
|  |    108 | lemma single_stateE: 
 | 
|  |    109 |   "state_not_singleton ==> !t. (!s::state. s = t) --> False"
 | 
|  |    110 | apply (unfold state_not_singleton_def)
 | 
|  |    111 | apply clarify
 | 
|  |    112 | apply (case_tac "ta = t")
 | 
|  |    113 | apply blast
 | 
|  |    114 | apply (blast dest: not_sym)
 | 
|  |    115 | done
 | 
|  |    116 | 
 | 
|  |    117 | declare peek_and_def [simp]
 | 
|  |    118 | 
 | 
|  |    119 | 
 | 
|  |    120 | subsection "validity"
 | 
|  |    121 | 
 | 
|  |    122 | lemma triple_valid_def2: 
 | 
|  |    123 |   "|=n:{P}.c.{Q} = (!Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s'))"
 | 
|  |    124 | apply (unfold triple_valid_def)
 | 
|  |    125 | apply auto
 | 
|  |    126 | done
 | 
|  |    127 | 
 | 
|  |    128 | lemma Body_triple_valid_0: "|=0:{P}. BODY pn .{Q}"
 | 
|  |    129 | apply (simp (no_asm) add: triple_valid_def2)
 | 
|  |    130 | apply clarsimp
 | 
|  |    131 | done
 | 
|  |    132 | 
 | 
|  |    133 | (* only ==> direction required *)
 | 
|  |    134 | lemma Body_triple_valid_Suc: "|=n:{P}. the (body pn) .{Q} = |=Suc n:{P}. BODY pn .{Q}"
 | 
|  |    135 | apply (simp (no_asm) add: triple_valid_def2)
 | 
|  |    136 | apply force
 | 
|  |    137 | done
 | 
|  |    138 | 
 | 
|  |    139 | lemma triple_valid_Suc [rule_format (no_asm)]: "|=Suc n:t --> |=n:t"
 | 
|  |    140 | apply (unfold triple_valid_def)
 | 
|  |    141 | apply (induct_tac t)
 | 
|  |    142 | apply simp
 | 
|  |    143 | apply (fast intro: evaln_Suc)
 | 
|  |    144 | done
 | 
|  |    145 | 
 | 
|  |    146 | lemma triples_valid_Suc: "||=Suc n:ts ==> ||=n:ts"
 | 
|  |    147 | apply (fast intro: triple_valid_Suc)
 | 
|  |    148 | done
 | 
|  |    149 | 
 | 
|  |    150 | 
 | 
|  |    151 | subsection "derived rules"
 | 
|  |    152 | 
 | 
|  |    153 | lemma conseq12: "[| G|-{P'}.c.{Q'}; !Z s. P Z s -->  
 | 
|  |    154 |                          (!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s') |]  
 | 
|  |    155 |        ==> G|-{P}.c.{Q}"
 | 
|  |    156 | apply (rule hoare_derivs.conseq)
 | 
|  |    157 | apply blast
 | 
|  |    158 | done
 | 
|  |    159 | 
 | 
|  |    160 | lemma conseq1: "[| G|-{P'}.c.{Q}; !Z s. P Z s --> P' Z s |] ==> G|-{P}.c.{Q}"
 | 
|  |    161 | apply (erule conseq12)
 | 
|  |    162 | apply fast
 | 
|  |    163 | done
 | 
|  |    164 | 
 | 
|  |    165 | lemma conseq2: "[| G|-{P}.c.{Q'}; !Z s. Q' Z s --> Q Z s |] ==> G|-{P}.c.{Q}"
 | 
|  |    166 | apply (erule conseq12)
 | 
|  |    167 | apply fast
 | 
|  |    168 | done
 | 
|  |    169 | 
 | 
|  |    170 | lemma Body1: "[| G Un (%p. {P p}.      BODY p  .{Q p})`Procs   
 | 
|  |    171 |           ||- (%p. {P p}. the (body p) .{Q p})`Procs;  
 | 
|  |    172 |     pn:Procs |] ==> G|-{P pn}. BODY pn .{Q pn}"
 | 
|  |    173 | apply (drule hoare_derivs.Body)
 | 
|  |    174 | apply (erule hoare_derivs.weaken)
 | 
|  |    175 | apply fast
 | 
|  |    176 | done
 | 
|  |    177 | 
 | 
|  |    178 | lemma BodyN: "(insert ({P}. BODY pn .{Q}) G) |-{P}. the (body pn) .{Q} ==>  
 | 
|  |    179 |   G|-{P}. BODY pn .{Q}"
 | 
|  |    180 | apply (rule Body1)
 | 
|  |    181 | apply (rule_tac [2] singletonI)
 | 
|  |    182 | apply clarsimp
 | 
|  |    183 | done
 | 
|  |    184 | 
 | 
|  |    185 | lemma escape: "[| !Z s. P Z s --> G|-{%Z s'. s'=s}.c.{%Z'. Q Z} |] ==> G|-{P}.c.{Q}"
 | 
|  |    186 | apply (rule hoare_derivs.conseq)
 | 
|  |    187 | apply fast
 | 
|  |    188 | done
 | 
|  |    189 | 
 | 
|  |    190 | lemma constant: "[| C ==> G|-{P}.c.{Q} |] ==> G|-{%Z s. P Z s & C}.c.{Q}"
 | 
|  |    191 | apply (rule hoare_derivs.conseq)
 | 
|  |    192 | apply fast
 | 
|  |    193 | done
 | 
|  |    194 | 
 | 
|  |    195 | lemma LoopF: "G|-{%Z s. P Z s & ~b s}.WHILE b DO c.{P}"
 | 
|  |    196 | apply (rule hoare_derivs.Loop [THEN conseq2])
 | 
|  |    197 | apply  (simp_all (no_asm))
 | 
|  |    198 | apply (rule hoare_derivs.conseq)
 | 
|  |    199 | apply fast
 | 
|  |    200 | done
 | 
|  |    201 | 
 | 
|  |    202 | (*
 | 
|  |    203 | Goal "[| G'||-ts; G' <= G |] ==> G||-ts"
 | 
|  |    204 | by (etac hoare_derivs.cut 1);
 | 
|  |    205 | by (etac hoare_derivs.asm 1);
 | 
|  |    206 | qed "thin";
 | 
|  |    207 | *)
 | 
|  |    208 | lemma thin [rule_format]: "G'||-ts ==> !G. G' <= G --> G||-ts"
 | 
|  |    209 | apply (erule hoare_derivs.induct)
 | 
|  |    210 | apply                (tactic {* ALLGOALS (EVERY'[Clarify_tac, REPEAT o smp_tac 1]) *})
 | 
|  |    211 | apply (rule hoare_derivs.empty)
 | 
|  |    212 | apply               (erule (1) hoare_derivs.insert)
 | 
|  |    213 | apply              (fast intro: hoare_derivs.asm)
 | 
|  |    214 | apply             (fast intro: hoare_derivs.cut)
 | 
|  |    215 | apply            (fast intro: hoare_derivs.weaken)
 | 
|  |    216 | apply           (rule hoare_derivs.conseq, intro strip, tactic "smp_tac 2 1", clarify, tactic "smp_tac 1 1",rule exI, rule exI, erule (1) conjI)
 | 
|  |    217 | prefer 7
 | 
|  |    218 | apply          (rule_tac hoare_derivs.Body, drule_tac spec, erule_tac mp, fast)
 | 
|  |    219 | apply         (tactic {* ALLGOALS (resolve_tac ((funpow 5 tl) (thms "hoare_derivs.intros")) THEN_ALL_NEW CLASET' fast_tac) *})
 | 
|  |    220 | done
 | 
|  |    221 | 
 | 
|  |    222 | lemma weak_Body: "G|-{P}. the (body pn) .{Q} ==> G|-{P}. BODY pn .{Q}"
 | 
|  |    223 | apply (rule BodyN)
 | 
|  |    224 | apply (erule thin)
 | 
|  |    225 | apply auto
 | 
|  |    226 | done
 | 
|  |    227 | 
 | 
|  |    228 | lemma derivs_insertD: "G||-insert t ts ==> G|-t & G||-ts"
 | 
|  |    229 | apply (fast intro: hoare_derivs.weaken)
 | 
|  |    230 | done
 | 
|  |    231 | 
 | 
|  |    232 | lemma finite_pointwise [rule_format (no_asm)]: "[| finite U;  
 | 
|  |    233 |   !p. G |-     {P' p}.c0 p.{Q' p}       --> G |-     {P p}.c0 p.{Q p} |] ==>  
 | 
|  |    234 |       G||-(%p. {P' p}.c0 p.{Q' p}) ` U --> G||-(%p. {P p}.c0 p.{Q p}) ` U"
 | 
|  |    235 | apply (erule finite_induct)
 | 
|  |    236 | apply simp
 | 
|  |    237 | apply clarsimp
 | 
|  |    238 | apply (drule derivs_insertD)
 | 
|  |    239 | apply (rule hoare_derivs.insert)
 | 
|  |    240 | apply  auto
 | 
|  |    241 | done
 | 
|  |    242 | 
 | 
|  |    243 | 
 | 
|  |    244 | subsection "soundness"
 | 
|  |    245 | 
 | 
|  |    246 | lemma Loop_sound_lemma: 
 | 
|  |    247 |  "G|={P &> b}. c .{P} ==>  
 | 
|  |    248 |   G|={P}. WHILE b DO c .{P &> (Not o b)}"
 | 
|  |    249 | apply (unfold hoare_valids_def)
 | 
|  |    250 | apply (simp (no_asm_use) add: triple_valid_def2)
 | 
|  |    251 | apply (rule allI)
 | 
|  |    252 | apply (subgoal_tac "!d s s'. <d,s> -n-> s' --> d = WHILE b DO c --> ||=n:G --> (!Z. P Z s --> P Z s' & ~b s') ")
 | 
|  |    253 | apply  (erule thin_rl, fast)
 | 
|  |    254 | apply ((rule allI)+, rule impI)
 | 
|  |    255 | apply (erule evaln.induct)
 | 
|  |    256 | apply (simp_all (no_asm))
 | 
|  |    257 | apply fast
 | 
|  |    258 | apply fast
 | 
|  |    259 | done
 | 
|  |    260 | 
 | 
|  |    261 | lemma Body_sound_lemma: 
 | 
|  |    262 |    "[| G Un (%pn. {P pn}.      BODY pn  .{Q pn})`Procs  
 | 
|  |    263 |          ||=(%pn. {P pn}. the (body pn) .{Q pn})`Procs |] ==>  
 | 
|  |    264 |         G||=(%pn. {P pn}.      BODY pn  .{Q pn})`Procs"
 | 
|  |    265 | apply (unfold hoare_valids_def)
 | 
|  |    266 | apply (rule allI)
 | 
|  |    267 | apply (induct_tac n)
 | 
|  |    268 | apply  (fast intro: Body_triple_valid_0)
 | 
|  |    269 | apply clarsimp
 | 
|  |    270 | apply (drule triples_valid_Suc)
 | 
|  |    271 | apply (erule (1) notE impE)
 | 
|  |    272 | apply (simp add: ball_Un)
 | 
|  |    273 | apply (drule spec, erule impE, erule conjI, assumption)
 | 
|  |    274 | apply (fast intro!: Body_triple_valid_Suc [THEN iffD1])
 | 
|  |    275 | done
 | 
|  |    276 | 
 | 
|  |    277 | lemma hoare_sound: "G||-ts ==> G||=ts"
 | 
|  |    278 | apply (erule hoare_derivs.induct)
 | 
|  |    279 | apply              (tactic {* TRYALL (eresolve_tac [thm "Loop_sound_lemma", thm "Body_sound_lemma"] THEN_ALL_NEW atac) *})
 | 
|  |    280 | apply            (unfold hoare_valids_def)
 | 
|  |    281 | apply            blast
 | 
|  |    282 | apply           blast
 | 
|  |    283 | apply          (blast) (* asm *)
 | 
|  |    284 | apply         (blast) (* cut *)
 | 
|  |    285 | apply        (blast) (* weaken *)
 | 
|  |    286 | apply       (tactic {* ALLGOALS (EVERY'[REPEAT o thin_tac "?x : hoare_derivs", SIMPSET' simp_tac, CLASET' clarify_tac, REPEAT o smp_tac 1]) *})
 | 
|  |    287 | apply       (simp_all (no_asm_use) add: triple_valid_def2)
 | 
|  |    288 | apply       (intro strip, tactic "smp_tac 2 1", blast) (* conseq *)
 | 
|  |    289 | apply      (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *}) (* Skip, Ass, Local *)
 | 
|  |    290 | prefer 3 apply   (force) (* Call *)
 | 
|  |    291 | apply  (erule_tac [2] evaln_elim_cases) (* If *)
 | 
|  |    292 | apply   blast+
 | 
|  |    293 | done
 | 
|  |    294 | 
 | 
|  |    295 | 
 | 
|  |    296 | section "completeness"
 | 
|  |    297 | 
 | 
|  |    298 | (* Both versions *)
 | 
|  |    299 | 
 | 
|  |    300 | (*unused*)
 | 
|  |    301 | lemma MGT_alternI: "G|-MGT c ==>  
 | 
|  |    302 |   G|-{%Z s0. !s1. <c,s0> -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1}"
 | 
|  |    303 | apply (unfold MGT_def)
 | 
|  |    304 | apply (erule conseq12)
 | 
|  |    305 | apply auto
 | 
|  |    306 | done
 | 
|  |    307 | 
 | 
|  |    308 | (* requires com_det *)
 | 
|  |    309 | lemma MGT_alternD: "state_not_singleton ==>  
 | 
|  |    310 |   G|-{%Z s0. !s1. <c,s0> -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1} ==> G|-MGT c"
 | 
|  |    311 | apply (unfold MGT_def)
 | 
|  |    312 | apply (erule conseq12)
 | 
|  |    313 | apply auto
 | 
|  |    314 | apply (case_tac "? t. <c,?s> -c-> t")
 | 
|  |    315 | apply  (fast elim: com_det)
 | 
|  |    316 | apply clarsimp
 | 
|  |    317 | apply (drule single_stateE)
 | 
|  |    318 | apply blast
 | 
|  |    319 | done
 | 
|  |    320 | 
 | 
|  |    321 | lemma MGF_complete: 
 | 
|  |    322 |  "{}|-(MGT c::state triple) ==> {}|={P}.c.{Q} ==> {}|-{P}.c.{Q::state assn}"
 | 
|  |    323 | apply (unfold MGT_def)
 | 
|  |    324 | apply (erule conseq12)
 | 
|  |    325 | apply (clarsimp simp add: hoare_valids_def eval_eq triple_valid_def2)
 | 
|  |    326 | done
 | 
|  |    327 | 
 | 
|  |    328 | declare WTs_elim_cases [elim!]
 | 
|  |    329 | declare not_None_eq [iff]
 | 
|  |    330 | (* requires com_det, escape (i.e. hoare_derivs.conseq) *)
 | 
|  |    331 | lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton ==>  
 | 
|  |    332 |   !pn:dom body. G|-{=}.BODY pn.{->} ==> WT c --> G|-{=}.c.{->}"
 | 
|  |    333 | apply (induct_tac c)
 | 
|  |    334 | apply        (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *})
 | 
|  |    335 | prefer 7 apply        (fast intro: domI)
 | 
|  |    336 | apply (erule_tac [6] MGT_alternD)
 | 
|  |    337 | apply       (unfold MGT_def)
 | 
|  |    338 | apply       (drule_tac [7] bspec, erule_tac [7] domI)
 | 
|  |    339 | apply       (rule_tac [7] escape, tactic {* CLASIMPSET' clarsimp_tac 7 *},
 | 
|  |    340 |   rule_tac [7] P1 = "%Z' s. s= (setlocs Z newlocs) [Loc Arg ::= fun Z]" in hoare_derivs.Call [THEN conseq1], erule_tac [7] conseq12)
 | 
|  |    341 | apply        (erule_tac [!] thin_rl)
 | 
|  |    342 | apply (rule hoare_derivs.Skip [THEN conseq2])
 | 
|  |    343 | apply (rule_tac [2] hoare_derivs.Ass [THEN conseq1])
 | 
|  |    344 | apply        (rule_tac [3] escape, tactic {* CLASIMPSET' clarsimp_tac 3 *},
 | 
|  |    345 |   rule_tac [3] P1 = "%Z' s. s= (Z[Loc loc::=fun Z])" in hoare_derivs.Local [THEN conseq1],
 | 
|  |    346 |   erule_tac [3] conseq12)
 | 
|  |    347 | apply         (erule_tac [5] hoare_derivs.Comp, erule_tac [5] conseq12)
 | 
|  |    348 | apply         (tactic {* (rtac (thm "hoare_derivs.If") THEN_ALL_NEW etac (thm "conseq12")) 6 *})
 | 
|  |    349 | apply          (rule_tac [8] hoare_derivs.Loop [THEN conseq2], erule_tac [8] conseq12)
 | 
|  |    350 | apply           auto
 | 
|  |    351 | done
 | 
|  |    352 | 
 | 
|  |    353 | (* Version: nested single recursion *)
 | 
|  |    354 | 
 | 
|  |    355 | lemma nesting_lemma [rule_format]:
 | 
|  |    356 |   assumes "!!G ts. ts <= G ==> P G ts"
 | 
|  |    357 |     and "!!G pn. P (insert (mgt_call pn) G) {mgt(the(body pn))} ==> P G {mgt_call pn}"
 | 
|  |    358 |     and "!!G c. [| wt c; !pn:U. P G {mgt_call pn} |] ==> P G {mgt c}"
 | 
|  |    359 |     and "!!pn. pn : U ==> wt (the (body pn))"
 | 
|  |    360 |   shows "finite U ==> uG = mgt_call`U ==>  
 | 
|  |    361 |   !G. G <= uG --> n <= card uG --> card G = card uG - n --> (!c. wt c --> P G {mgt c})"
 | 
|  |    362 | apply (induct_tac n)
 | 
|  |    363 | apply  (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *})
 | 
|  |    364 | apply  (subgoal_tac "G = mgt_call ` U")
 | 
|  |    365 | prefer 2
 | 
|  |    366 | apply   (simp add: card_seteq finite_imageI)
 | 
|  |    367 | apply  simp
 | 
|  |    368 | apply  (erule prems(3-)) (*MGF_lemma1*)
 | 
|  |    369 | apply (rule ballI)
 | 
|  |    370 | apply  (rule prems) (*hoare_derivs.asm*)
 | 
|  |    371 | apply  fast
 | 
|  |    372 | apply (erule prems(3-)) (*MGF_lemma1*)
 | 
|  |    373 | apply (rule ballI)
 | 
|  |    374 | apply (case_tac "mgt_call pn : G")
 | 
|  |    375 | apply  (rule prems) (*hoare_derivs.asm*)
 | 
|  |    376 | apply  fast
 | 
|  |    377 | apply (rule prems(2-)) (*MGT_BodyN*)
 | 
|  |    378 | apply (drule spec, erule impE, erule_tac [2] impE, drule_tac [3] spec, erule_tac [3] mp)
 | 
|  |    379 | apply   (erule_tac [3] prems(4-))
 | 
|  |    380 | apply   fast
 | 
|  |    381 | apply (drule finite_subset)
 | 
|  |    382 | apply (erule finite_imageI)
 | 
|  |    383 | apply (simp (no_asm_simp))
 | 
|  |    384 | done
 | 
|  |    385 | 
 | 
|  |    386 | lemma MGT_BodyN: "insert ({=}.BODY pn.{->}) G|-{=}. the (body pn) .{->} ==>  
 | 
|  |    387 |   G|-{=}.BODY pn.{->}"
 | 
|  |    388 | apply (unfold MGT_def)
 | 
|  |    389 | apply (rule BodyN)
 | 
|  |    390 | apply (erule conseq2)
 | 
|  |    391 | apply force
 | 
|  |    392 | done
 | 
|  |    393 | 
 | 
|  |    394 | (* requires BodyN, com_det *)
 | 
|  |    395 | lemma MGF: "[| state_not_singleton; WT_bodies; WT c |] ==> {}|-MGT c"
 | 
|  |    396 | apply (rule_tac P = "%G ts. G||-ts" and U = "dom body" in nesting_lemma)
 | 
|  |    397 | apply (erule hoare_derivs.asm)
 | 
|  |    398 | apply (erule MGT_BodyN)
 | 
|  |    399 | apply (rule_tac [3] finite_dom_body)
 | 
|  |    400 | apply (erule MGF_lemma1)
 | 
|  |    401 | prefer 2 apply (assumption)
 | 
|  |    402 | apply       blast
 | 
|  |    403 | apply      clarsimp
 | 
|  |    404 | apply     (erule (1) WT_bodiesD)
 | 
|  |    405 | apply (rule_tac [3] le_refl)
 | 
|  |    406 | apply    auto
 | 
|  |    407 | done
 | 
|  |    408 | 
 | 
|  |    409 | 
 | 
|  |    410 | (* Version: simultaneous recursion in call rule *)
 | 
|  |    411 | 
 | 
|  |    412 | (* finiteness not really necessary here *)
 | 
|  |    413 | lemma MGT_Body: "[| G Un (%pn. {=}.      BODY pn  .{->})`Procs  
 | 
|  |    414 |                           ||-(%pn. {=}. the (body pn) .{->})`Procs;  
 | 
|  |    415 |   finite Procs |] ==>   G ||-(%pn. {=}.      BODY pn  .{->})`Procs"
 | 
|  |    416 | apply (unfold MGT_def)
 | 
|  |    417 | apply (rule hoare_derivs.Body)
 | 
|  |    418 | apply (erule finite_pointwise)
 | 
|  |    419 | prefer 2 apply (assumption)
 | 
|  |    420 | apply clarify
 | 
|  |    421 | apply (erule conseq2)
 | 
|  |    422 | apply auto
 | 
|  |    423 | done
 | 
|  |    424 | 
 | 
|  |    425 | (* requires empty, insert, com_det *)
 | 
|  |    426 | lemma MGF_lemma2_simult [rule_format (no_asm)]: "[| state_not_singleton; WT_bodies;  
 | 
|  |    427 |   F<=(%pn. {=}.the (body pn).{->})`dom body |] ==>  
 | 
|  |    428 |      (%pn. {=}.     BODY pn .{->})`dom body||-F"
 | 
|  |    429 | apply (frule finite_subset)
 | 
|  |    430 | apply (rule finite_dom_body [THEN finite_imageI])
 | 
|  |    431 | apply (rotate_tac 2)
 | 
|  |    432 | apply (tactic "make_imp_tac 1")
 | 
|  |    433 | apply (erule finite_induct)
 | 
|  |    434 | apply  (clarsimp intro!: hoare_derivs.empty)
 | 
|  |    435 | apply (clarsimp intro!: hoare_derivs.insert simp del: range_composition)
 | 
|  |    436 | apply (erule MGF_lemma1)
 | 
|  |    437 | prefer 2 apply  (fast dest: WT_bodiesD)
 | 
|  |    438 | apply clarsimp
 | 
|  |    439 | apply (rule hoare_derivs.asm)
 | 
|  |    440 | apply (fast intro: domI)
 | 
|  |    441 | done
 | 
|  |    442 | 
 | 
|  |    443 | (* requires Body, empty, insert, com_det *)
 | 
|  |    444 | lemma MGF': "[| state_not_singleton; WT_bodies; WT c |] ==> {}|-MGT c"
 | 
|  |    445 | apply (rule MGF_lemma1)
 | 
|  |    446 | apply assumption
 | 
|  |    447 | prefer 2 apply (assumption)
 | 
|  |    448 | apply clarsimp
 | 
|  |    449 | apply (subgoal_tac "{}||- (%pn. {=}. BODY pn .{->}) `dom body")
 | 
|  |    450 | apply (erule hoare_derivs.weaken)
 | 
|  |    451 | apply  (fast intro: domI)
 | 
|  |    452 | apply (rule finite_dom_body [THEN [2] MGT_Body])
 | 
|  |    453 | apply (simp (no_asm))
 | 
|  |    454 | apply (erule (1) MGF_lemma2_simult)
 | 
|  |    455 | apply (rule subset_refl)
 | 
|  |    456 | done
 | 
|  |    457 | 
 | 
|  |    458 | (* requires Body+empty+insert / BodyN, com_det *)
 | 
|  |    459 | lemmas hoare_complete = MGF' [THEN MGF_complete, standard]
 | 
|  |    460 | 
 | 
|  |    461 | 
 | 
|  |    462 | subsection "unused derived rules"
 | 
|  |    463 | 
 | 
|  |    464 | lemma falseE: "G|-{%Z s. False}.c.{Q}"
 | 
|  |    465 | apply (rule hoare_derivs.conseq)
 | 
|  |    466 | apply fast
 | 
|  |    467 | done
 | 
|  |    468 | 
 | 
|  |    469 | lemma trueI: "G|-{P}.c.{%Z s. True}"
 | 
|  |    470 | apply (rule hoare_derivs.conseq)
 | 
|  |    471 | apply (fast intro!: falseE)
 | 
|  |    472 | done
 | 
|  |    473 | 
 | 
|  |    474 | lemma disj: "[| G|-{P}.c.{Q}; G|-{P'}.c.{Q'} |]  
 | 
|  |    475 |         ==> G|-{%Z s. P Z s | P' Z s}.c.{%Z s. Q Z s | Q' Z s}"
 | 
|  |    476 | apply (rule hoare_derivs.conseq)
 | 
|  |    477 | apply (fast elim: conseq12)
 | 
|  |    478 | done (* analogue conj non-derivable *)
 | 
|  |    479 | 
 | 
|  |    480 | lemma hoare_SkipI: "(!Z s. P Z s --> Q Z s) ==> G|-{P}. SKIP .{Q}"
 | 
|  |    481 | apply (rule conseq12)
 | 
|  |    482 | apply (rule hoare_derivs.Skip)
 | 
|  |    483 | apply fast
 | 
|  |    484 | done
 | 
|  |    485 | 
 | 
|  |    486 | 
 | 
|  |    487 | subsection "useful derived rules"
 | 
|  |    488 | 
 | 
|  |    489 | lemma single_asm: "{t}|-t"
 | 
|  |    490 | apply (rule hoare_derivs.asm)
 | 
|  |    491 | apply (rule subset_refl)
 | 
|  |    492 | done
 | 
|  |    493 | 
 | 
|  |    494 | lemma export_s: "[| !!s'. G|-{%Z s. s'=s & P Z s}.c.{Q} |] ==> G|-{P}.c.{Q}"
 | 
|  |    495 | apply (rule hoare_derivs.conseq)
 | 
|  |    496 | apply auto
 | 
|  |    497 | done
 | 
|  |    498 | 
 | 
|  |    499 | 
 | 
|  |    500 | lemma weak_Local: "[| G|-{P}. c .{Q}; !k Z s. Q Z s --> Q Z (s[Loc Y::=k]) |] ==>  
 | 
|  |    501 |     G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{Q}"
 | 
|  |    502 | apply (rule export_s)
 | 
|  |    503 | apply (rule hoare_derivs.Local)
 | 
|  |    504 | apply (erule conseq2)
 | 
|  |    505 | apply (erule spec)
 | 
|  |    506 | done
 | 
|  |    507 | 
 | 
|  |    508 | (*
 | 
|  |    509 | Goal "!Q. G |-{%Z s. ~(? s'. <c,s> -c-> s')}. c .{Q}"
 | 
|  |    510 | by (induct_tac "c" 1);
 | 
|  |    511 | by Auto_tac;
 | 
|  |    512 | by (rtac conseq1 1);
 | 
|  |    513 | by (rtac hoare_derivs.Skip 1);
 | 
|  |    514 | force 1;
 | 
|  |    515 | by (rtac conseq1 1);
 | 
|  |    516 | by (rtac hoare_derivs.Ass 1);
 | 
|  |    517 | force 1;
 | 
|  |    518 | by (defer_tac 1);
 | 
|  |    519 | ###
 | 
|  |    520 | by (rtac hoare_derivs.Comp 1);
 | 
|  |    521 | by (dtac spec 2);
 | 
|  |    522 | by (dtac spec 2);
 | 
|  |    523 | by (assume_tac 2);
 | 
|  |    524 | by (etac conseq1 2);
 | 
|  |    525 | by (Clarsimp_tac 2);
 | 
|  |    526 | force 1;
 | 
|  |    527 | *)
 | 
| 8177 |    528 | 
 | 
|  |    529 | end
 |