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