src/HOL/Hoare/hoare_tac.ML
changeset 24475 a297ae4ff188
child 26300 03def556e26e
equal deleted inserted replaced
24474:33da394f0888 24475:a297ae4ff188
       
     1 (*  Title:      HOL/Hoare/hoare_tac.ML
       
     2     ID:         $Id$
       
     3     Author:     Leonor Prensa Nieto & Tobias Nipkow
       
     4     Copyright   1998 TUM
       
     5 
       
     6 Derivation of the proof rules and, most importantly, the VCG tactic.
       
     7 *)
       
     8 
       
     9 (*** The tactics ***)
       
    10 
       
    11 (*****************************************************************************)
       
    12 (** The function Mset makes the theorem                                     **)
       
    13 (** "?Mset <= {(x1,...,xn). ?P (x1,...,xn)} ==> ?Mset <= {s. ?P s}",        **)
       
    14 (** where (x1,...,xn) are the variables of the particular program we are    **)
       
    15 (** working on at the moment of the call                                    **)
       
    16 (*****************************************************************************)
       
    17 
       
    18 local open HOLogic in
       
    19 
       
    20 (** maps (%x1 ... xn. t) to [x1,...,xn] **)
       
    21 fun abs2list (Const ("split",_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
       
    22   | abs2list (Abs(x,T,t)) = [Free (x, T)]
       
    23   | abs2list _ = [];
       
    24 
       
    25 (** maps {(x1,...,xn). t} to [x1,...,xn] **)
       
    26 fun mk_vars (Const ("Collect",_) $ T) = abs2list T
       
    27   | mk_vars _ = [];
       
    28 
       
    29 (** abstraction of body over a tuple formed from a list of free variables. 
       
    30 Types are also built **)
       
    31 fun mk_abstupleC []     body = absfree ("x", unitT, body)
       
    32   | mk_abstupleC (v::w) body = let val (n,T) = dest_Free v
       
    33                                in if w=[] then absfree (n, T, body)
       
    34         else let val z  = mk_abstupleC w body;
       
    35                  val T2 = case z of Abs(_,T,_) => T
       
    36                         | Const (_, Type (_,[_, Type (_,[T,_])])) $ _ => T;
       
    37        in Const ("split", (T --> T2 --> boolT) --> mk_prodT (T,T2) --> boolT) 
       
    38           $ absfree (n, T, z) end end;
       
    39 
       
    40 (** maps [x1,...,xn] to (x1,...,xn) and types**)
       
    41 fun mk_bodyC []      = HOLogic.unit
       
    42   | mk_bodyC (x::xs) = if xs=[] then x 
       
    43                else let val (n, T) = dest_Free x ;
       
    44                         val z = mk_bodyC xs;
       
    45                         val T2 = case z of Free(_, T) => T
       
    46                                          | Const ("Pair", Type ("fun", [_, Type
       
    47                                             ("fun", [_, T])])) $ _ $ _ => T;
       
    48                  in Const ("Pair", [T, T2] ---> mk_prodT (T, T2)) $ x $ z end;
       
    49 
       
    50 (** maps a goal of the form:
       
    51         1. [| P |] ==> VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn]**) 
       
    52 fun get_vars thm = let  val c = Logic.unprotect (concl_of (thm));
       
    53                         val d = Logic.strip_assums_concl c;
       
    54                         val Const _ $ pre $ _ $ _ = dest_Trueprop d;
       
    55       in mk_vars pre end;
       
    56 
       
    57 
       
    58 (** Makes Collect with type **)
       
    59 fun mk_CollectC trm = let val T as Type ("fun",[t,_]) = fastype_of trm 
       
    60                       in Collect_const t $ trm end;
       
    61 
       
    62 fun inclt ty = Const (@{const_name HOL.less_eq}, [ty,ty] ---> boolT);
       
    63 
       
    64 (** Makes "Mset <= t" **)
       
    65 fun Mset_incl t = let val MsetT = fastype_of t 
       
    66                  in mk_Trueprop ((inclt MsetT) $ Free ("Mset", MsetT) $ t) end;
       
    67 
       
    68 
       
    69 fun Mset thm = let val vars = get_vars(thm);
       
    70                    val varsT = fastype_of (mk_bodyC vars);
       
    71                    val big_Collect = mk_CollectC (mk_abstupleC vars 
       
    72                          (Free ("P",varsT --> boolT) $ mk_bodyC vars));
       
    73                    val small_Collect = mk_CollectC (Abs("x",varsT,
       
    74                            Free ("P",varsT --> boolT) $ Bound 0));
       
    75                    val impl = implies $ (Mset_incl big_Collect) $ 
       
    76                                           (Mset_incl small_Collect);
       
    77    in Goal.prove (ProofContext.init (Thm.theory_of_thm thm)) ["Mset", "P"] [] impl (K (CLASET' blast_tac 1)) end;
       
    78 
       
    79 end;
       
    80 
       
    81 
       
    82 (*****************************************************************************)
       
    83 (** Simplifying:                                                            **)
       
    84 (** Some useful lemmata, lists and simplification tactics to control which  **)
       
    85 (** theorems are used to simplify at each moment, so that the original      **)
       
    86 (** input does not suffer any unexpected transformation                     **)
       
    87 (*****************************************************************************)
       
    88 
       
    89 Goal "-(Collect b) = {x. ~(b x)}";
       
    90 by (Fast_tac 1);
       
    91 qed "Compl_Collect";
       
    92 
       
    93 
       
    94 (**Simp_tacs**)
       
    95 
       
    96 val before_set2pred_simp_tac =
       
    97   (simp_tac (HOL_basic_ss addsimps [Collect_conj_eq RS sym,Compl_Collect]));
       
    98 
       
    99 val split_simp_tac = (simp_tac (HOL_basic_ss addsimps [split_conv]));
       
   100 
       
   101 (*****************************************************************************)
       
   102 (** set2pred transforms sets inclusion into predicates implication,         **)
       
   103 (** maintaining the original variable names.                                **)
       
   104 (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1"              **)
       
   105 (** Subgoals containing intersections (A Int B) or complement sets (-A)     **)
       
   106 (** are first simplified by "before_set2pred_simp_tac", that returns only   **)
       
   107 (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily           **)
       
   108 (** transformed.                                                            **)
       
   109 (** This transformation may solve very easy subgoals due to a ligth         **)
       
   110 (** simplification done by (split_all_tac)                                  **)
       
   111 (*****************************************************************************)
       
   112 
       
   113 fun set2pred i thm =
       
   114   let val var_names = map (fst o dest_Free) (get_vars thm) in
       
   115     ((before_set2pred_simp_tac i) THEN_MAYBE
       
   116      (EVERY [rtac subsetI i, 
       
   117              rtac CollectI i,
       
   118              dtac CollectD i,
       
   119              (TRY(split_all_tac i)) THEN_MAYBE
       
   120              ((rename_params_tac var_names i) THEN
       
   121               (full_simp_tac (HOL_basic_ss addsimps [split_conv]) i)) ])) thm
       
   122   end;
       
   123 
       
   124 (*****************************************************************************)
       
   125 (** BasicSimpTac is called to simplify all verification conditions. It does **)
       
   126 (** a light simplification by applying "mem_Collect_eq", then it calls      **)
       
   127 (** MaxSimpTac, which solves subgoals of the form "A <= A",                 **)
       
   128 (** and transforms any other into predicates, applying then                 **)
       
   129 (** the tactic chosen by the user, which may solve the subgoal completely.  **)
       
   130 (*****************************************************************************)
       
   131 
       
   132 fun MaxSimpTac tac = FIRST'[rtac subset_refl, set2pred THEN_MAYBE' tac];
       
   133 
       
   134 fun BasicSimpTac tac =
       
   135   simp_tac
       
   136     (HOL_basic_ss addsimps [mem_Collect_eq,split_conv] addsimprocs [record_simproc])
       
   137   THEN_MAYBE' MaxSimpTac tac;
       
   138 
       
   139 (** HoareRuleTac **)
       
   140 
       
   141 fun WlpTac Mlem tac i =
       
   142   rtac @{thm SeqRule} i THEN  HoareRuleTac Mlem tac false (i+1)
       
   143 and HoareRuleTac Mlem tac pre_cond i st = st |>
       
   144         (*abstraction over st prevents looping*)
       
   145     ( (WlpTac Mlem tac i THEN HoareRuleTac Mlem tac pre_cond i)
       
   146       ORELSE
       
   147       (FIRST[rtac @{thm SkipRule} i,
       
   148              EVERY[rtac @{thm BasicRule} i,
       
   149                    rtac Mlem i,
       
   150                    split_simp_tac i],
       
   151              EVERY[rtac @{thm CondRule} i,
       
   152                    HoareRuleTac Mlem tac false (i+2),
       
   153                    HoareRuleTac Mlem tac false (i+1)],
       
   154              EVERY[rtac @{thm WhileRule} i,
       
   155                    BasicSimpTac tac (i+2),
       
   156                    HoareRuleTac Mlem tac true (i+1)] ] 
       
   157        THEN (if pre_cond then (BasicSimpTac tac i) else (rtac subset_refl i)) ));
       
   158 
       
   159 
       
   160 (** tac:(int -> tactic) is the tactic the user chooses to solve or simplify **)
       
   161 (** the final verification conditions                                       **)
       
   162  
       
   163 fun hoare_tac tac i thm =
       
   164   let val Mlem = Mset(thm)
       
   165   in SELECT_GOAL(EVERY[HoareRuleTac Mlem tac true 1]) i thm end;