| author | wenzelm | 
| Thu, 15 Jan 2015 14:01:26 +0100 | |
| changeset 59370 | b13ff987c559 | 
| parent 55661 | ec14796cd140 | 
| child 60754 | 02924903a6fd | 
| permissions | -rw-r--r-- | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Hoare/hoare_tac.ML  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Leonor Prensa Nieto & Tobias Nipkow  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
Derivation of the proof rules and, most importantly, the VCG tactic.  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 55660 | 7  | 
signature HOARE =  | 
8  | 
sig  | 
|
| 55661 | 9  | 
val hoare_rule_tac: Proof.context -> term list * thm -> (int -> tactic) -> bool ->  | 
10  | 
int -> tactic  | 
|
| 55660 | 11  | 
val hoare_tac: Proof.context -> (int -> tactic) -> int -> tactic  | 
12  | 
end;  | 
|
13  | 
||
14  | 
structure Hoare: HOARE =  | 
|
15  | 
struct  | 
|
| 41449 | 16  | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
(*** The tactics ***)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
(*****************************************************************************)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
(** The function Mset makes the theorem **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
(** "?Mset <= {(x1,...,xn). ?P (x1,...,xn)} ==> ?Mset <= {s. ?P s}",        **)
 | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
(** where (x1,...,xn) are the variables of the particular program we are **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
(** working on at the moment of the call **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
(*****************************************************************************)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
|
| 41449 | 26  | 
local  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
(** maps (%x1 ... xn. t) to [x1,...,xn] **)  | 
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
51717 
diff
changeset
 | 
29  | 
fun abs2list (Const (@{const_name case_prod}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
 | 
| 55659 | 30  | 
| abs2list (Abs (x, T, _)) = [Free (x, T)]  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
| abs2list _ = [];  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
(** maps {(x1,...,xn). t} to [x1,...,xn] **)
 | 
| 37677 | 34  | 
fun mk_vars (Const (@{const_name Collect},_) $ T) = abs2list T
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
| mk_vars _ = [];  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
37  | 
(** abstraction of body over a tuple formed from a list of free variables.  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
Types are also built **)  | 
| 44241 | 39  | 
fun mk_abstupleC [] body = absfree ("x", HOLogic.unitT) body
 | 
40  | 
| mk_abstupleC [v] body = absfree (dest_Free v) body  | 
|
41  | 
| mk_abstupleC (v :: w) body =  | 
|
42  | 
let  | 
|
43  | 
val (x, T) = dest_Free v;  | 
|
44  | 
val z = mk_abstupleC w body;  | 
|
45  | 
val T2 =  | 
|
46  | 
(case z of  | 
|
47  | 
Abs (_, T, _) => T  | 
|
48  | 
| Const (_, Type (_, [_, Type (_, [T, _])])) $ _ => T);  | 
|
49  | 
in  | 
|
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
51717 
diff
changeset
 | 
50  | 
        Const (@{const_name case_prod},
 | 
| 44241 | 51  | 
(T --> T2 --> HOLogic.boolT) --> HOLogic.mk_prodT (T, T2) --> HOLogic.boolT) $  | 
52  | 
absfree (x, T) z  | 
|
53  | 
end;  | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
(** maps [x1,...,xn] to (x1,...,xn) and types**)  | 
| 55661 | 56  | 
fun mk_bodyC [] = HOLogic.unit  | 
57  | 
| mk_bodyC [x] = x  | 
|
58  | 
| mk_bodyC (x :: xs) =  | 
|
59  | 
let  | 
|
60  | 
val (_, T) = dest_Free x;  | 
|
61  | 
val z = mk_bodyC xs;  | 
|
62  | 
val T2 =  | 
|
63  | 
(case z of  | 
|
64  | 
Free (_, T) => T  | 
|
65  | 
          | Const (@{const_name Pair}, Type ("fun", [_, Type ("fun", [_, T])])) $ _ $ _ => T);
 | 
|
66  | 
     in Const (@{const_name Pair}, [T, T2] ---> HOLogic.mk_prodT (T, T2)) $ x $ z end;
 | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
68  | 
(** maps a subgoal of the form:  | 
| 55661 | 69  | 
    VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn]
 | 
70  | 
**)  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
71  | 
fun get_vars c =  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
72  | 
let  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
73  | 
val d = Logic.strip_assums_concl c;  | 
| 41449 | 74  | 
val Const _ $ pre $ _ $ _ = HOLogic.dest_Trueprop d;  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
75  | 
in mk_vars pre end;  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
|
| 55661 | 77  | 
fun mk_CollectC tm =  | 
78  | 
  let val T as Type ("fun",[t,_]) = fastype_of tm;
 | 
|
79  | 
in HOLogic.Collect_const t $ tm end;  | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
|
| 41449 | 81  | 
fun inclt ty = Const (@{const_name Orderings.less_eq}, [ty,ty] ---> HOLogic.boolT);
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
|
| 41449 | 83  | 
in  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
85  | 
fun Mset ctxt prop =  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
86  | 
let  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
87  | 
    val [(Mset, _), (P, _)] = Variable.variant_frees ctxt [] [("Mset", ()), ("P", ())];
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
89  | 
val vars = get_vars prop;  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
90  | 
val varsT = fastype_of (mk_bodyC vars);  | 
| 55661 | 91  | 
val big_Collect =  | 
92  | 
mk_CollectC (mk_abstupleC vars (Free (P, varsT --> HOLogic.boolT) $ mk_bodyC vars));  | 
|
93  | 
val small_Collect =  | 
|
94  | 
      mk_CollectC (Abs ("x", varsT, Free (P, varsT --> HOLogic.boolT) $ Bound 0));
 | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
95  | 
|
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
96  | 
val MsetT = fastype_of big_Collect;  | 
| 41449 | 97  | 
fun Mset_incl t = HOLogic.mk_Trueprop (inclt MsetT $ Free (Mset, MsetT) $ t);  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
98  | 
val impl = Logic.mk_implies (Mset_incl big_Collect, Mset_incl small_Collect);  | 
| 42793 | 99  | 
val th = Goal.prove ctxt [Mset, P] [] impl (fn _ => blast_tac ctxt 1);  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
100  | 
in (vars, th) end;  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
end;  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
(*****************************************************************************)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
(** Simplifying: **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
(** Some useful lemmata, lists and simplification tactics to control which **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
(** theorems are used to simplify at each moment, so that the original **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
(** input does not suffer any unexpected transformation **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
(*****************************************************************************)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
(**Simp_tacs**)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
114  | 
fun before_set2pred_simp_tac ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
115  | 
  simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Collect_conj_eq RS sym, @{thm Compl_Collect}]);
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
117  | 
fun split_simp_tac ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
118  | 
  simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]);
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
(*****************************************************************************)  | 
| 55661 | 121  | 
(** set_to_pred_tac transforms sets inclusion into predicates implication, **)  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
(** maintaining the original variable names. **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
(** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1"              **)
 | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
(** Subgoals containing intersections (A Int B) or complement sets (-A) **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
(** are first simplified by "before_set2pred_simp_tac", that returns only **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
(** subgoals of the form "{x. P x} <= {x. Q x}", which are easily           **)
 | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
(** transformed. **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
(** This transformation may solve very easy subgoals due to a ligth **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
(** simplification done by (split_all_tac) **)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
(*****************************************************************************)  | 
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
|
| 55661 | 132  | 
fun set_to_pred_tac ctxt var_names = SUBGOAL (fn (_, i) =>  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
133  | 
before_set2pred_simp_tac ctxt i THEN_MAYBE  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
134  | 
EVERY [  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
135  | 
rtac subsetI i,  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
136  | 
rtac CollectI i,  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
137  | 
dtac CollectD i,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
138  | 
TRY (split_all_tac ctxt i) THEN_MAYBE  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
139  | 
(rename_tac var_names i THEN  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
140  | 
      full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i)]);
 | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
|
| 55661 | 142  | 
(*******************************************************************************)  | 
143  | 
(** basic_simp_tac is called to simplify all verification conditions. It does **)  | 
|
144  | 
(** a light simplification by applying "mem_Collect_eq", then it calls **)  | 
|
145  | 
(** max_simp_tac, which solves subgoals of the form "A <= A", **)  | 
|
146  | 
(** and transforms any other into predicates, applying then **)  | 
|
147  | 
(** the tactic chosen by the user, which may solve the subgoal completely. **)  | 
|
148  | 
(*******************************************************************************)  | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
|
| 55661 | 150  | 
fun max_simp_tac ctxt var_names tac =  | 
151  | 
FIRST' [rtac subset_refl, set_to_pred_tac ctxt var_names THEN_MAYBE' tac];  | 
|
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
|
| 55661 | 153  | 
fun basic_simp_tac ctxt var_names tac =  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
154  | 
simp_tac  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
155  | 
(put_simpset HOL_basic_ss ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
156  | 
      addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [Record.simproc])
 | 
| 55661 | 157  | 
THEN_MAYBE' max_simp_tac ctxt var_names tac;  | 
| 
24475
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
|
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
160  | 
(** hoare_rule_tac **)  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
161  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
162  | 
fun hoare_rule_tac ctxt (vars, Mlem) tac =  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
163  | 
let  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
164  | 
val var_names = map (fst o dest_Free) vars;  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
165  | 
fun wlp_tac i =  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
166  | 
      rtac @{thm SeqRule} i THEN rule_tac false (i + 1)
 | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
167  | 
and rule_tac pre_cond i st = st |> (*abstraction over st prevents looping*)  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
168  | 
((wlp_tac i THEN rule_tac pre_cond i)  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
169  | 
ORELSE  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
170  | 
(FIRST [  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
171  | 
          rtac @{thm SkipRule} i,
 | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
172  | 
          rtac @{thm AbortRule} i,
 | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
173  | 
EVERY [  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
174  | 
            rtac @{thm BasicRule} i,
 | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
175  | 
rtac Mlem i,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
176  | 
split_simp_tac ctxt i],  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
177  | 
EVERY [  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
178  | 
            rtac @{thm CondRule} i,
 | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
179  | 
rule_tac false (i + 2),  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
180  | 
rule_tac false (i + 1)],  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
181  | 
EVERY [  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
182  | 
            rtac @{thm WhileRule} i,
 | 
| 55661 | 183  | 
basic_simp_tac ctxt var_names tac (i + 2),  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
184  | 
rule_tac true (i + 1)]]  | 
| 55661 | 185  | 
THEN (if pre_cond then basic_simp_tac ctxt var_names tac i else rtac subset_refl i)));  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
186  | 
in rule_tac end;  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
187  | 
|
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
188  | 
|
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
189  | 
(** tac is the tactic the user chooses to solve or simplify **)  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
190  | 
(** the final verification conditions **)  | 
| 
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
191  | 
|
| 55659 | 192  | 
fun hoare_tac ctxt tac = SUBGOAL (fn (goal, i) =>  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
193  | 
SELECT_GOAL (hoare_rule_tac ctxt (Mset ctxt goal) tac true 1) i);  | 
| 
28457
 
25669513fd4c
major cleanup of hoare_tac.ML: just one copy for Hoare.thy and HoareAbort.thy (only 1 line different), refrain from inspecting the main goal, proper context;
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
194  | 
|
| 55660 | 195  | 
end;  | 
196  |