| 
1476
 | 
     1  | 
(*  Title:      HOL/Hoare/Hoare.thy
  | 
| 
1335
 | 
     2  | 
    ID:         $Id$
  | 
| 
5646
 | 
     3  | 
    Author:     Leonor Prensa Nieto & Tobias Nipkow
  | 
| 
 | 
     4  | 
    Copyright   1998 TUM
  | 
| 
1335
 | 
     5  | 
  | 
| 
 | 
     6  | 
Sugared semantic embedding of Hoare logic.
  | 
| 
5646
 | 
     7  | 
Strictly speaking a shallow embedding (as implemented by Norbert Galm
  | 
| 
 | 
     8  | 
following Mike Gordon) would suffice. Maybe the datatype com comes in useful
  | 
| 
 | 
     9  | 
later.
  | 
| 
1335
 | 
    10  | 
*)
  | 
| 
 | 
    11  | 
  | 
| 
16417
 | 
    12  | 
theory Hoare  imports Main
  | 
| 
 | 
    13  | 
uses ("hoare.ML") begin
 | 
| 
1335
 | 
    14  | 
  | 
| 
 | 
    15  | 
types
  | 
| 
13682
 | 
    16  | 
    'a bexp = "'a set"
  | 
| 
 | 
    17  | 
    'a assn = "'a set"
  | 
| 
5646
 | 
    18  | 
  | 
| 
 | 
    19  | 
datatype
  | 
| 
13682
 | 
    20  | 
 'a com = Basic "'a \<Rightarrow> 'a"         
  | 
| 
 | 
    21  | 
   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
 | 
| 
 | 
    22  | 
   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
 | 
| 
 | 
    23  | 
   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
 | 
| 
5646
 | 
    24  | 
  
  | 
| 
 | 
    25  | 
syntax
  | 
| 
13682
 | 
    26  | 
  "@assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
 | 
| 
 | 
    27  | 
  "@annskip" :: "'a com"                    ("SKIP")
 | 
| 
5646
 | 
    28  | 
  | 
| 
 | 
    29  | 
translations
  | 
| 
 | 
    30  | 
            "SKIP" == "Basic id"
  | 
| 
 | 
    31  | 
  | 
| 
13682
 | 
    32  | 
types 'a sem = "'a => 'a => bool"
  | 
| 
5646
 | 
    33  | 
  | 
| 
13682
 | 
    34  | 
consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
  | 
| 
5646
 | 
    35  | 
primrec
  | 
| 
 | 
    36  | 
"iter 0 b S = (%s s'. s ~: b & (s=s'))"
  | 
| 
 | 
    37  | 
"iter (Suc n) b S = (%s s'. s : b & (? s''. S s s'' & iter n b S s'' s'))"
  | 
| 
 | 
    38  | 
  | 
| 
13682
 | 
    39  | 
consts Sem :: "'a com => 'a sem"
  | 
| 
5646
 | 
    40  | 
primrec
  | 
| 
 | 
    41  | 
"Sem(Basic f) s s' = (s' = f s)"
  | 
| 
 | 
    42  | 
"Sem(c1;c2) s s' = (? s''. Sem c1 s s'' & Sem c2 s'' s')"
  | 
| 
 | 
    43  | 
"Sem(IF b THEN c1 ELSE c2 FI) s s' = ((s  : b --> Sem c1 s s') &
  | 
| 
 | 
    44  | 
                                      (s ~: b --> Sem c2 s s'))"
  | 
| 
 | 
    45  | 
"Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
  | 
| 
 | 
    46  | 
  | 
| 
13682
 | 
    47  | 
constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
  | 
| 
5646
 | 
    48  | 
  "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
  | 
| 
5007
 | 
    49  | 
  | 
| 
 | 
    50  | 
  | 
| 
1335
 | 
    51  | 
syntax
  | 
| 
13764
 | 
    52  | 
 "@hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
  | 
| 
13737
 | 
    53  | 
                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
 | 
| 
5646
 | 
    54  | 
syntax ("" output)
 | 
| 
13682
 | 
    55  | 
 "@hoare"      :: "['a assn,'a com,'a assn] => bool"
  | 
| 
13737
 | 
    56  | 
                 ("{_} // _ // {_}" [0,55,0] 50)
 | 
| 
1335
 | 
    57  | 
  | 
| 
5646
 | 
    58  | 
(** parse translations **)
  | 
| 
1335
 | 
    59  | 
  | 
| 
13682
 | 
    60  | 
ML{*
 | 
| 
13764
 | 
    61  | 
  | 
| 
 | 
    62  | 
local
  | 
| 
13857
 | 
    63  | 
  | 
| 
13764
 | 
    64  | 
fun abs((a,T),body) =
  | 
| 
 | 
    65  | 
  let val a = absfree(a, dummyT, body)
  | 
| 
 | 
    66  | 
  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
  | 
| 
 | 
    67  | 
in
  | 
| 
1335
 | 
    68  | 
  | 
| 
13764
 | 
    69  | 
fun mk_abstuple [x] body = abs (x, body)
  | 
| 
 | 
    70  | 
  | mk_abstuple (x::xs) body =
  | 
| 
 | 
    71  | 
      Syntax.const "split" $ abs (x, mk_abstuple xs body);
  | 
| 
1335
 | 
    72  | 
  | 
| 
13857
 | 
    73  | 
fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
  | 
| 
13764
 | 
    74  | 
  | mk_fbody a e ((b,_)::xs) =
  | 
| 
13857
 | 
    75  | 
      Syntax.const "Pair" $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
  | 
| 
13764
 | 
    76  | 
  | 
| 
 | 
    77  | 
fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
  | 
| 
 | 
    78  | 
end
  | 
| 
13682
 | 
    79  | 
*}
  | 
| 
1335
 | 
    80  | 
  | 
| 
5646
 | 
    81  | 
(* bexp_tr & assn_tr *)
  | 
| 
9397
 | 
    82  | 
(*all meta-variables for bexp except for TRUE are translated as if they
  | 
| 
5646
 | 
    83  | 
  were boolean expressions*)
  | 
| 
13682
 | 
    84  | 
ML{*
 | 
| 
5646
 | 
    85  | 
fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"
 | 
| 
 | 
    86  | 
  | bexp_tr b xs = Syntax.const "Collect" $ mk_abstuple xs b;
  | 
| 
 | 
    87  | 
  
  | 
| 
 | 
    88  | 
fun assn_tr r xs = Syntax.const "Collect" $ mk_abstuple xs r;
  | 
| 
13682
 | 
    89  | 
*}
  | 
| 
5646
 | 
    90  | 
(* com_tr *)
  | 
| 
13682
 | 
    91  | 
ML{*
 | 
| 
13764
 | 
    92  | 
fun com_tr (Const("@assign",_) $ Free (a,_) $ e) xs =
 | 
| 
 | 
    93  | 
      Syntax.const "Basic" $ mk_fexp a e xs
  | 
| 
5646
 | 
    94  | 
  | com_tr (Const ("Basic",_) $ f) xs = Syntax.const "Basic" $ f
 | 
| 
13764
 | 
    95  | 
  | com_tr (Const ("Seq",_) $ c1 $ c2) xs =
 | 
| 
 | 
    96  | 
      Syntax.const "Seq" $ com_tr c1 xs $ com_tr c2 xs
  | 
| 
 | 
    97  | 
  | com_tr (Const ("Cond",_) $ b $ c1 $ c2) xs =
 | 
| 
 | 
    98  | 
      Syntax.const "Cond" $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
  | 
| 
 | 
    99  | 
  | com_tr (Const ("While",_) $ b $ I $ c) xs =
 | 
| 
 | 
   100  | 
      Syntax.const "While" $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
  | 
| 
 | 
   101  | 
  | com_tr t _ = t (* if t is just a Free/Var *)
  | 
| 
13682
 | 
   102  | 
*}
  | 
| 
5646
 | 
   103  | 
  | 
| 
17781
 | 
   104  | 
(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
  | 
| 
13682
 | 
   105  | 
ML{*
 | 
| 
13764
 | 
   106  | 
local
  | 
| 
 | 
   107  | 
  | 
| 
 | 
   108  | 
fun var_tr(Free(a,_)) = (a,Bound 0) (* Bound 0 = dummy term *)
  | 
| 
 | 
   109  | 
  | var_tr(Const ("_constrain", _) $ (Free (a,_)) $ T) = (a,T);
 | 
| 
5646
 | 
   110  | 
  | 
| 
13764
 | 
   111  | 
fun vars_tr (Const ("_idts", _) $ idt $ vars) = var_tr idt :: vars_tr vars
 | 
| 
 | 
   112  | 
  | vars_tr t = [var_tr t]
  | 
| 
 | 
   113  | 
  | 
| 
 | 
   114  | 
in
  | 
| 
5646
 | 
   115  | 
fun hoare_vars_tr [vars, pre, prg, post] =
  | 
| 
 | 
   116  | 
      let val xs = vars_tr vars
  | 
| 
 | 
   117  | 
      in Syntax.const "Valid" $
  | 
| 
13764
 | 
   118  | 
         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
  | 
| 
5646
 | 
   119  | 
      end
  | 
| 
 | 
   120  | 
  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
 | 
| 
13764
 | 
   121  | 
end
  | 
| 
13682
 | 
   122  | 
*}
  | 
| 
5646
 | 
   123  | 
  | 
| 
13682
 | 
   124  | 
parse_translation {* [("@hoare_vars", hoare_vars_tr)] *}
 | 
| 
1335
 | 
   125  | 
  | 
| 
 | 
   126  | 
  | 
| 
5646
 | 
   127  | 
(*****************************************************************************)
  | 
| 
 | 
   128  | 
  | 
| 
 | 
   129  | 
(*** print translations ***)
  | 
| 
13682
 | 
   130  | 
ML{*
 | 
| 
5646
 | 
   131  | 
fun dest_abstuple (Const ("split",_) $ (Abs(v,_, body))) =
 | 
| 
 | 
   132  | 
                            subst_bound (Syntax.free v, dest_abstuple body)
  | 
| 
 | 
   133  | 
  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
  | 
| 
 | 
   134  | 
  | dest_abstuple trm = trm;
  | 
| 
1335
 | 
   135  | 
  | 
| 
5646
 | 
   136  | 
fun abs2list (Const ("split",_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
 | 
| 
 | 
   137  | 
  | abs2list (Abs(x,T,t)) = [Free (x, T)]
  | 
| 
 | 
   138  | 
  | abs2list _ = [];
  | 
| 
1335
 | 
   139  | 
  | 
| 
5646
 | 
   140  | 
fun mk_ts (Const ("split",_) $ (Abs(x,_,t))) = mk_ts t
 | 
| 
 | 
   141  | 
  | mk_ts (Abs(x,_,t)) = mk_ts t
  | 
| 
 | 
   142  | 
  | mk_ts (Const ("Pair",_) $ a $ b) = a::(mk_ts b)
 | 
| 
 | 
   143  | 
  | mk_ts t = [t];
  | 
| 
1335
 | 
   144  | 
  | 
| 
5646
 | 
   145  | 
fun mk_vts (Const ("split",_) $ (Abs(x,_,t))) = 
 | 
| 
 | 
   146  | 
           ((Syntax.free x)::(abs2list t), mk_ts t)
  | 
| 
 | 
   147  | 
  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
  | 
| 
 | 
   148  | 
  | mk_vts t = raise Match;
  | 
| 
 | 
   149  | 
  
  | 
| 
 | 
   150  | 
fun find_ch [] i xs = (false, (Syntax.free "not_ch",Syntax.free "not_ch" ))
  | 
| 
 | 
   151  | 
  | find_ch ((v,t)::vts) i xs = if t=(Bound i) then find_ch vts (i-1) xs
  | 
| 
 | 
   152  | 
              else (true, (v, subst_bounds (xs,t)));
  | 
| 
 | 
   153  | 
  
  | 
| 
 | 
   154  | 
fun is_f (Const ("split",_) $ (Abs(x,_,t))) = true
 | 
| 
 | 
   155  | 
  | is_f (Abs(x,_,t)) = true
  | 
| 
 | 
   156  | 
  | is_f t = false;
  | 
| 
13682
 | 
   157  | 
*}
  | 
| 
 | 
   158  | 
  | 
| 
5646
 | 
   159  | 
(* assn_tr' & bexp_tr'*)
  | 
| 
13682
 | 
   160  | 
ML{*  
 | 
| 
5646
 | 
   161  | 
fun assn_tr' (Const ("Collect",_) $ T) = dest_abstuple T
 | 
| 
 | 
   162  | 
  | assn_tr' (Const ("op Int",_) $ (Const ("Collect",_) $ T1) $ 
 | 
| 
 | 
   163  | 
                                   (Const ("Collect",_) $ T2)) =  
 | 
| 
 | 
   164  | 
            Syntax.const "op Int" $ dest_abstuple T1 $ dest_abstuple T2
  | 
| 
 | 
   165  | 
  | assn_tr' t = t;
  | 
| 
1335
 | 
   166  | 
  | 
| 
5646
 | 
   167  | 
fun bexp_tr' (Const ("Collect",_) $ T) = dest_abstuple T 
 | 
| 
 | 
   168  | 
  | bexp_tr' t = t;
  | 
| 
13682
 | 
   169  | 
*}
  | 
| 
5646
 | 
   170  | 
  | 
| 
 | 
   171  | 
(*com_tr' *)
  | 
| 
13682
 | 
   172  | 
ML{*
 | 
| 
5646
 | 
   173  | 
fun mk_assign f =
  | 
| 
 | 
   174  | 
  let val (vs, ts) = mk_vts f;
  | 
| 
 | 
   175  | 
      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
  | 
| 
 | 
   176  | 
  in if ch then Syntax.const "@assign" $ fst(which) $ snd(which)
  | 
| 
 | 
   177  | 
     else Syntax.const "@skip" end;
  | 
| 
1335
 | 
   178  | 
  | 
| 
5646
 | 
   179  | 
fun com_tr' (Const ("Basic",_) $ f) = if is_f f then mk_assign f
 | 
| 
 | 
   180  | 
                                           else Syntax.const "Basic" $ f
  | 
| 
 | 
   181  | 
  | com_tr' (Const ("Seq",_) $ c1 $ c2) = Syntax.const "Seq" $
 | 
| 
 | 
   182  | 
                                                 com_tr' c1 $ com_tr' c2
  | 
| 
 | 
   183  | 
  | com_tr' (Const ("Cond",_) $ b $ c1 $ c2) = Syntax.const "Cond" $
 | 
| 
 | 
   184  | 
                                           bexp_tr' b $ com_tr' c1 $ com_tr' c2
  | 
| 
 | 
   185  | 
  | com_tr' (Const ("While",_) $ b $ I $ c) = Syntax.const "While" $
 | 
| 
 | 
   186  | 
                                               bexp_tr' b $ assn_tr' I $ com_tr' c
  | 
| 
 | 
   187  | 
  | com_tr' t = t;
  | 
| 
1335
 | 
   188  | 
  | 
| 
 | 
   189  | 
  | 
| 
5646
 | 
   190  | 
fun spec_tr' [p, c, q] =
  | 
| 
 | 
   191  | 
  Syntax.const "@hoare" $ assn_tr' p $ com_tr' c $ assn_tr' q
  | 
| 
13682
 | 
   192  | 
*}
  | 
| 
 | 
   193  | 
  | 
| 
 | 
   194  | 
print_translation {* [("Valid", spec_tr')] *}
 | 
| 
 | 
   195  | 
  | 
| 
13857
 | 
   196  | 
lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
  | 
| 
 | 
   197  | 
by (auto simp:Valid_def)
  | 
| 
 | 
   198  | 
  | 
| 
 | 
   199  | 
lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
 | 
| 
 | 
   200  | 
by (auto simp:Valid_def)
  | 
| 
 | 
   201  | 
  | 
| 
 | 
   202  | 
lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
  | 
| 
 | 
   203  | 
by (auto simp:Valid_def)
  | 
| 
 | 
   204  | 
  | 
| 
 | 
   205  | 
lemma CondRule:
  | 
| 
 | 
   206  | 
 "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
 | 
| 
 | 
   207  | 
  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
  | 
| 
 | 
   208  | 
by (auto simp:Valid_def)
  | 
| 
 | 
   209  | 
  | 
| 
 | 
   210  | 
lemma iter_aux: "! s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
  | 
| 
 | 
   211  | 
       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)";
  | 
| 
 | 
   212  | 
apply(induct n)
  | 
| 
 | 
   213  | 
 apply clarsimp
  | 
| 
 | 
   214  | 
apply(simp (no_asm_use))
  | 
| 
 | 
   215  | 
apply blast
  | 
| 
 | 
   216  | 
done
  | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
lemma WhileRule:
  | 
| 
 | 
   219  | 
 "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
  | 
| 
 | 
   220  | 
apply (clarsimp simp:Valid_def)
  | 
| 
 | 
   221  | 
apply(drule iter_aux)
  | 
| 
 | 
   222  | 
  prefer 2 apply assumption
  | 
| 
 | 
   223  | 
 apply blast
  | 
| 
 | 
   224  | 
apply blast
  | 
| 
 | 
   225  | 
done
  | 
| 
 | 
   226  | 
  | 
| 
 | 
   227  | 
  | 
| 
13696
 | 
   228  | 
use "hoare.ML"
  | 
| 
13682
 | 
   229  | 
  | 
| 
 | 
   230  | 
method_setup vcg = {*
 | 
| 
21588
 | 
   231  | 
  Method.no_args (Method.SIMPLE_METHOD' (hoare_tac (K all_tac))) *}
  | 
| 
13682
 | 
   232  | 
  "verification condition generator"
  | 
| 
 | 
   233  | 
  | 
| 
 | 
   234  | 
method_setup vcg_simp = {*
 | 
| 
 | 
   235  | 
  Method.ctxt_args (fn ctxt =>
  | 
| 
21588
 | 
   236  | 
    Method.SIMPLE_METHOD' (hoare_tac (asm_full_simp_tac (local_simpset_of ctxt)))) *}
  | 
| 
13682
 | 
   237  | 
  "verification condition generator plus simplification"
  | 
| 
 | 
   238  | 
  | 
| 
 | 
   239  | 
end
  |