author | wenzelm |
Fri, 21 Feb 2014 20:37:13 +0100 | |
changeset 55660 | f0f895716a8b |
parent 55659 | 4089f6e98ab9 |
child 55661 | ec14796cd140 |
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 |
|
9 |
val hoare_rule_tac: Proof.context -> term list * thm -> (int -> tactic) -> bool -> int -> tactic |
|
10 |
val hoare_tac: Proof.context -> (int -> tactic) -> int -> tactic |
|
11 |
end; |
|
12 |
||
13 |
structure Hoare: HOARE = |
|
14 |
struct |
|
41449 | 15 |
|
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
16 |
(*** The tactics ***) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
17 |
|
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 |
(** 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
|
20 |
(** "?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
|
21 |
(** 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
|
22 |
(** 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
|
23 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
24 |
|
41449 | 25 |
local |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
26 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
27 |
(** maps (%x1 ... xn. t) to [x1,...,xn] **) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
51717
diff
changeset
|
28 |
fun abs2list (Const (@{const_name case_prod}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t |
55659 | 29 |
| 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
|
30 |
| abs2list _ = []; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
31 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
32 |
(** maps {(x1,...,xn). t} to [x1,...,xn] **) |
37677 | 33 |
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
|
34 |
| mk_vars _ = []; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
35 |
|
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
|
36 |
(** 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
|
37 |
Types are also built **) |
44241 | 38 |
fun mk_abstupleC [] body = absfree ("x", HOLogic.unitT) body |
39 |
| mk_abstupleC [v] body = absfree (dest_Free v) body |
|
40 |
| mk_abstupleC (v :: w) body = |
|
41 |
let |
|
42 |
val (x, T) = dest_Free v; |
|
43 |
val z = mk_abstupleC w body; |
|
44 |
val T2 = |
|
45 |
(case z of |
|
46 |
Abs (_, T, _) => T |
|
47 |
| Const (_, Type (_, [_, Type (_, [T, _])])) $ _ => T); |
|
48 |
in |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
51717
diff
changeset
|
49 |
Const (@{const_name case_prod}, |
44241 | 50 |
(T --> T2 --> HOLogic.boolT) --> HOLogic.mk_prodT (T, T2) --> HOLogic.boolT) $ |
51 |
absfree (x, T) z |
|
52 |
end; |
|
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
53 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
54 |
(** maps [x1,...,xn] to (x1,...,xn) and types**) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
55 |
fun mk_bodyC [] = HOLogic.unit |
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
|
56 |
| mk_bodyC (x::xs) = if xs=[] then x |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
57 |
else let val (n, T) = dest_Free x ; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
58 |
val z = mk_bodyC xs; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
59 |
val T2 = case z of Free(_, T) => T |
37391 | 60 |
| Const (@{const_name Pair}, Type ("fun", [_, Type |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
61 |
("fun", [_, T])])) $ _ $ _ => T; |
41449 | 62 |
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
|
63 |
|
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
|
64 |
(** maps a subgoal of the form: |
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
|
65 |
VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn]**) |
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
|
66 |
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
|
67 |
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
|
68 |
val d = Logic.strip_assums_concl c; |
41449 | 69 |
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
|
70 |
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
|
71 |
|
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
|
72 |
fun mk_CollectC trm = |
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 |
let val T as Type ("fun",[t,_]) = fastype_of trm |
41449 | 74 |
in HOLogic.Collect_const t $ trm end; |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
75 |
|
41449 | 76 |
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
|
77 |
|
41449 | 78 |
in |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
79 |
|
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
|
80 |
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
|
81 |
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
|
82 |
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
|
83 |
|
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
|
84 |
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
|
85 |
val varsT = fastype_of (mk_bodyC vars); |
41449 | 86 |
val big_Collect = mk_CollectC (mk_abstupleC vars (Free (P, varsT --> HOLogic.boolT) $ mk_bodyC vars)); |
87 |
val small_Collect = 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
|
88 |
|
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 MsetT = fastype_of big_Collect; |
41449 | 90 |
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
|
91 |
val impl = Logic.mk_implies (Mset_incl big_Collect, Mset_incl small_Collect); |
42793 | 92 |
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
|
93 |
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
|
94 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
95 |
end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
96 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
97 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
98 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
99 |
(** Simplifying: **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
100 |
(** 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
|
101 |
(** 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
|
102 |
(** 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
|
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 |
(**Simp_tacs**) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
106 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
107 |
fun before_set2pred_simp_tac ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
108 |
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
|
109 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
110 |
fun split_simp_tac ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
111 |
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
|
112 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
113 |
(*****************************************************************************) |
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
|
114 |
(** set2pred_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
|
115 |
(** 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
|
116 |
(** 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
|
117 |
(** 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
|
118 |
(** 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
|
119 |
(** 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
|
120 |
(** transformed. **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
121 |
(** 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
|
122 |
(** 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
|
123 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
124 |
|
55659 | 125 |
fun set2pred_tac ctxt var_names = SUBGOAL (fn (_, i) => |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
126 |
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
|
127 |
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
|
128 |
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
|
129 |
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
|
130 |
dtac CollectD i, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
131 |
TRY (split_all_tac ctxt i) THEN_MAYBE |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
132 |
(rename_tac var_names i THEN |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
133 |
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
|
134 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
135 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
136 |
(** BasicSimpTac is called to simplify all verification conditions. It does **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
137 |
(** a light simplification by applying "mem_Collect_eq", then it calls **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
138 |
(** MaxSimpTac, which solves subgoals of the form "A <= A", **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
139 |
(** and transforms any other into predicates, applying then **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
140 |
(** the tactic chosen by the user, which may solve the subgoal completely. **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
141 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
142 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
143 |
fun MaxSimpTac ctxt var_names tac = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
144 |
FIRST'[rtac subset_refl, set2pred_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
|
145 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
146 |
fun BasicSimpTac 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
|
147 |
simp_tac |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
148 |
(put_simpset HOL_basic_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
149 |
addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [Record.simproc]) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
150 |
THEN_MAYBE' MaxSimpTac 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
|
151 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
152 |
|
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
|
153 |
(** 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
|
154 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
155 |
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
|
156 |
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
|
157 |
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
|
158 |
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
|
159 |
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
|
160 |
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
|
161 |
((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
|
162 |
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
|
163 |
(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
|
164 |
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
|
165 |
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
|
166 |
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
|
167 |
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
|
168 |
rtac Mlem i, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
169 |
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
|
170 |
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
|
171 |
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
|
172 |
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
|
173 |
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
|
174 |
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
|
175 |
rtac @{thm WhileRule} i, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
176 |
BasicSimpTac 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
|
177 |
rule_tac true (i + 1)]] |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44241
diff
changeset
|
178 |
THEN (if pre_cond then BasicSimpTac 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
|
179 |
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
|
180 |
|
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 |
|
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 |
(** 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
|
183 |
(** 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
|
184 |
|
55659 | 185 |
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
|
186 |
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
|
187 |
|
55660 | 188 |
end; |
189 |