src/HOL/Hoare/HoareAbort.thy
changeset 24470 41c81e23c08d
parent 21588 cd0dc678a205
child 27244 af0a44372d1f
equal deleted inserted replaced
24469:01fd2863d7c8 24470:41c81e23c08d
     5 
     5 
     6 Like Hoare.thy, but with an Abort statement for modelling run time errors.
     6 Like Hoare.thy, but with an Abort statement for modelling run time errors.
     7 *)
     7 *)
     8 
     8 
     9 theory HoareAbort  imports Main
     9 theory HoareAbort  imports Main
    10 uses ("hoareAbort.ML") begin
    10 begin
    11 
    11 
    12 types
    12 types
    13     'a bexp = "'a set"
    13     'a bexp = "'a set"
    14     'a assn = "'a set"
    14     'a assn = "'a set"
    15 
    15 
   233 done
   233 done
   234 
   234 
   235 lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
   235 lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
   236 by(auto simp:Valid_def)
   236 by(auto simp:Valid_def)
   237 
   237 
   238 use "hoareAbort.ML"
   238 
       
   239 subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
       
   240 
       
   241 ML {*
       
   242 (*** The tactics ***)
       
   243 
       
   244 (*****************************************************************************)
       
   245 (** The function Mset makes the theorem                                     **)
       
   246 (** "?Mset <= {(x1,...,xn). ?P (x1,...,xn)} ==> ?Mset <= {s. ?P s}",        **)
       
   247 (** where (x1,...,xn) are the variables of the particular program we are    **)
       
   248 (** working on at the moment of the call                                    **)
       
   249 (*****************************************************************************)
       
   250 
       
   251 local open HOLogic in
       
   252 
       
   253 (** maps (%x1 ... xn. t) to [x1,...,xn] **)
       
   254 fun abs2list (Const ("split",_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
       
   255   | abs2list (Abs(x,T,t)) = [Free (x, T)]
       
   256   | abs2list _ = [];
       
   257 
       
   258 (** maps {(x1,...,xn). t} to [x1,...,xn] **)
       
   259 fun mk_vars (Const ("Collect",_) $ T) = abs2list T
       
   260   | mk_vars _ = [];
       
   261 
       
   262 (** abstraction of body over a tuple formed from a list of free variables. 
       
   263 Types are also built **)
       
   264 fun mk_abstupleC []     body = absfree ("x", unitT, body)
       
   265   | mk_abstupleC (v::w) body = let val (n,T) = dest_Free v
       
   266                                in if w=[] then absfree (n, T, body)
       
   267         else let val z  = mk_abstupleC w body;
       
   268                  val T2 = case z of Abs(_,T,_) => T
       
   269                         | Const (_, Type (_,[_, Type (_,[T,_])])) $ _ => T;
       
   270        in Const ("split", (T --> T2 --> boolT) --> mk_prodT (T,T2) --> boolT) 
       
   271           $ absfree (n, T, z) end end;
       
   272 
       
   273 (** maps [x1,...,xn] to (x1,...,xn) and types**)
       
   274 fun mk_bodyC []      = HOLogic.unit
       
   275   | mk_bodyC (x::xs) = if xs=[] then x 
       
   276                else let val (n, T) = dest_Free x ;
       
   277                         val z = mk_bodyC xs;
       
   278                         val T2 = case z of Free(_, T) => T
       
   279                                          | Const ("Pair", Type ("fun", [_, Type
       
   280                                             ("fun", [_, T])])) $ _ $ _ => T;
       
   281                  in Const ("Pair", [T, T2] ---> mk_prodT (T, T2)) $ x $ z end;
       
   282 
       
   283 (** maps a goal of the form:
       
   284         1. [| P |] ==> VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn]**) 
       
   285 fun get_vars thm = let  val c = Logic.unprotect (concl_of (thm));
       
   286                         val d = Logic.strip_assums_concl c;
       
   287                         val Const _ $ pre $ _ $ _ = dest_Trueprop d;
       
   288       in mk_vars pre end;
       
   289 
       
   290 
       
   291 (** Makes Collect with type **)
       
   292 fun mk_CollectC trm = let val T as Type ("fun",[t,_]) = fastype_of trm 
       
   293                       in Collect_const t $ trm end;
       
   294 
       
   295 fun inclt ty = Const (@{const_name HOL.less_eq}, [ty,ty] ---> boolT);
       
   296 
       
   297 (** Makes "Mset <= t" **)
       
   298 fun Mset_incl t = let val MsetT = fastype_of t 
       
   299                  in mk_Trueprop ((inclt MsetT) $ Free ("Mset", MsetT) $ t) end;
       
   300 
       
   301 
       
   302 fun Mset thm = let val vars = get_vars(thm);
       
   303                    val varsT = fastype_of (mk_bodyC vars);
       
   304                    val big_Collect = mk_CollectC (mk_abstupleC vars 
       
   305                          (Free ("P",varsT --> boolT) $ mk_bodyC vars));
       
   306                    val small_Collect = mk_CollectC (Abs("x",varsT,
       
   307                            Free ("P",varsT --> boolT) $ Bound 0));
       
   308                    val impl = implies $ (Mset_incl big_Collect) $ 
       
   309                                           (Mset_incl small_Collect);
       
   310    in Goal.prove (ProofContext.init (Thm.theory_of_thm thm)) ["Mset", "P"] [] impl (K (CLASET' blast_tac 1)) end;
       
   311 
       
   312 end;
       
   313 *}
       
   314 
       
   315 (*****************************************************************************)
       
   316 (** Simplifying:                                                            **)
       
   317 (** Some useful lemmata, lists and simplification tactics to control which  **)
       
   318 (** theorems are used to simplify at each moment, so that the original      **)
       
   319 (** input does not suffer any unexpected transformation                     **)
       
   320 (*****************************************************************************)
       
   321 
       
   322 lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
       
   323   by blast
       
   324 
       
   325 
       
   326 ML {*
       
   327 (**Simp_tacs**)
       
   328 
       
   329 val before_set2pred_simp_tac =
       
   330   (simp_tac (HOL_basic_ss addsimps [@{thm Collect_conj_eq} RS sym, @{thm Compl_Collect}]));
       
   331 
       
   332 val split_simp_tac = (simp_tac (HOL_basic_ss addsimps [split_conv]));
       
   333 
       
   334 (*****************************************************************************)
       
   335 (** set2pred transforms sets inclusion into predicates implication,         **)
       
   336 (** maintaining the original variable names.                                **)
       
   337 (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1"              **)
       
   338 (** Subgoals containing intersections (A Int B) or complement sets (-A)     **)
       
   339 (** are first simplified by "before_set2pred_simp_tac", that returns only   **)
       
   340 (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily           **)
       
   341 (** transformed.                                                            **)
       
   342 (** This transformation may solve very easy subgoals due to a ligth         **)
       
   343 (** simplification done by (split_all_tac)                                  **)
       
   344 (*****************************************************************************)
       
   345 
       
   346 fun set2pred i thm =
       
   347   let val var_names = map (fst o dest_Free) (get_vars thm) in
       
   348     ((before_set2pred_simp_tac i) THEN_MAYBE
       
   349       (EVERY [rtac subsetI i, 
       
   350               rtac CollectI i,
       
   351               dtac CollectD i,
       
   352               (TRY(split_all_tac i)) THEN_MAYBE
       
   353               ((rename_params_tac var_names i) THEN
       
   354                (full_simp_tac (HOL_basic_ss addsimps [split_conv]) i)) ])) thm
       
   355   end;
       
   356 
       
   357 (*****************************************************************************)
       
   358 (** BasicSimpTac is called to simplify all verification conditions. It does **)
       
   359 (** a light simplification by applying "mem_Collect_eq", then it calls      **)
       
   360 (** MaxSimpTac, which solves subgoals of the form "A <= A",                 **)
       
   361 (** and transforms any other into predicates, applying then                 **)
       
   362 (** the tactic chosen by the user, which may solve the subgoal completely.  **)
       
   363 (*****************************************************************************)
       
   364 
       
   365 fun MaxSimpTac tac = FIRST'[rtac subset_refl, set2pred THEN_MAYBE' tac];
       
   366 
       
   367 fun BasicSimpTac tac =
       
   368   simp_tac
       
   369     (HOL_basic_ss addsimps [mem_Collect_eq,split_conv] addsimprocs [record_simproc])
       
   370   THEN_MAYBE' MaxSimpTac tac;
       
   371 
       
   372 (** HoareRuleTac **)
       
   373 
       
   374 fun WlpTac Mlem tac i =
       
   375   rtac @{thm SeqRule} i THEN  HoareRuleTac Mlem tac false (i+1)
       
   376 and HoareRuleTac Mlem tac pre_cond i st = st |>
       
   377         (*abstraction over st prevents looping*)
       
   378     ( (WlpTac Mlem tac i THEN HoareRuleTac Mlem tac pre_cond i)
       
   379       ORELSE
       
   380       (FIRST[rtac @{thm SkipRule} i,
       
   381              rtac @{thm AbortRule} i,
       
   382              EVERY[rtac @{thm BasicRule} i,
       
   383                    rtac Mlem i,
       
   384                    split_simp_tac i],
       
   385              EVERY[rtac @{thm CondRule} i,
       
   386                    HoareRuleTac Mlem tac false (i+2),
       
   387                    HoareRuleTac Mlem tac false (i+1)],
       
   388              EVERY[rtac @{thm WhileRule} i,
       
   389                    BasicSimpTac tac (i+2),
       
   390                    HoareRuleTac Mlem tac true (i+1)] ] 
       
   391        THEN (if pre_cond then (BasicSimpTac tac i) else rtac subset_refl i) ));
       
   392 
       
   393 
       
   394 (** tac:(int -> tactic) is the tactic the user chooses to solve or simplify **)
       
   395 (** the final verification conditions                                       **)
       
   396  
       
   397 fun hoare_tac tac i thm =
       
   398   let val Mlem = Mset(thm)
       
   399   in SELECT_GOAL(EVERY[HoareRuleTac Mlem tac true 1]) i thm end;
       
   400 *}
   239 
   401 
   240 method_setup vcg = {*
   402 method_setup vcg = {*
   241   Method.no_args (Method.SIMPLE_METHOD' (hoare_tac (K all_tac))) *}
   403   Method.no_args (Method.SIMPLE_METHOD' (hoare_tac (K all_tac))) *}
   242   "verification condition generator"
   404   "verification condition generator"
   243 
   405