author | wenzelm |
Tue, 03 Mar 2009 14:07:43 +0100 | |
changeset 30211 | 556d1810cdad |
parent 28457 | 25669513fd4c |
child 32149 | ef59550a55d3 |
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 |
ID: $Id$ |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
3 |
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
|
4 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
5 |
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
|
6 |
*) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
7 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
8 |
(*** The tactics ***) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
9 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
10 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
11 |
(** 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
|
12 |
(** "?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
|
13 |
(** 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
|
14 |
(** 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
|
15 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
16 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
17 |
local open HOLogic in |
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 |
(** maps (%x1 ... xn. t) to [x1,...,xn] **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
20 |
fun abs2list (Const ("split",_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
21 |
| abs2list (Abs(x,T,t)) = [Free (x, T)] |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
22 |
| abs2list _ = []; |
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 |
(** maps {(x1,...,xn). t} to [x1,...,xn] **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
25 |
fun mk_vars (Const ("Collect",_) $ T) = abs2list T |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
26 |
| mk_vars _ = []; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
27 |
|
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
|
28 |
(** 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
|
29 |
Types are also built **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
30 |
fun mk_abstupleC [] body = absfree ("x", unitT, body) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
31 |
| mk_abstupleC (v::w) body = let val (n,T) = dest_Free v |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
32 |
in if w=[] then absfree (n, T, body) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
33 |
else let val z = mk_abstupleC w body; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
34 |
val T2 = case z of Abs(_,T,_) => T |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
35 |
| Const (_, Type (_,[_, Type (_,[T,_])])) $ _ => 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
|
36 |
in Const ("split", (T --> T2 --> boolT) --> mk_prodT (T,T2) --> boolT) |
24475
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
37 |
$ absfree (n, T, z) end end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
38 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
39 |
(** 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
|
40 |
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
|
41 |
| 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
|
42 |
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
|
43 |
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
|
44 |
val T2 = case z of Free(_, T) => T |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
45 |
| Const ("Pair", Type ("fun", [_, Type |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
46 |
("fun", [_, T])])) $ _ $ _ => T; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
47 |
in Const ("Pair", [T, T2] ---> mk_prodT (T, T2)) $ x $ z end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
48 |
|
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
|
49 |
(** 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
|
50 |
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
|
51 |
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
|
52 |
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
|
53 |
val d = Logic.strip_assums_concl 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
|
54 |
val Const _ $ pre $ _ $ _ = dest_Trueprop d; |
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
|
55 |
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
|
56 |
|
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
|
57 |
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
|
58 |
let val T as Type ("fun",[t,_]) = fastype_of 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
|
59 |
in 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
|
60 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
61 |
fun inclt ty = Const (@{const_name HOL.less_eq}, [ty,ty] ---> boolT); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
62 |
|
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 |
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
|
65 |
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
|
66 |
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
|
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 |
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
|
69 |
val varsT = fastype_of (mk_bodyC 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
|
70 |
val big_Collect = mk_CollectC (mk_abstupleC vars (Free (P, varsT --> boolT) $ mk_bodyC 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
|
71 |
val small_Collect = mk_CollectC (Abs ("x", varsT, Free (P, varsT --> boolT) $ Bound 0)); |
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 |
|
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 MsetT = fastype_of big_Collect; |
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
|
74 |
fun Mset_incl t = mk_Trueprop (inclt MsetT $ Free (Mset, MsetT) $ t); |
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 |
val impl = Logic.mk_implies (Mset_incl big_Collect, Mset_incl small_Collect); |
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
|
76 |
val th = Goal.prove ctxt [Mset, P] [] impl (fn _ => blast_tac (local_claset_of ctxt) 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
|
77 |
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
|
78 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
79 |
end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
80 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
81 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
82 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
83 |
(** Simplifying: **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
84 |
(** 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
|
85 |
(** 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
|
86 |
(** 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
|
87 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
88 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
89 |
(**Simp_tacs**) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
90 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
91 |
val before_set2pred_simp_tac = |
26300 | 92 |
(simp_tac (HOL_basic_ss 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
|
93 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
94 |
val split_simp_tac = (simp_tac (HOL_basic_ss addsimps [split_conv])); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
95 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
96 |
(*****************************************************************************) |
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
|
97 |
(** 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
|
98 |
(** 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
|
99 |
(** 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
|
100 |
(** 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
|
101 |
(** 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
|
102 |
(** 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
|
103 |
(** transformed. **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
104 |
(** 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
|
105 |
(** 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
|
106 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
107 |
|
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
|
108 |
fun set2pred_tac var_names = SUBGOAL (fn (goal, 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
|
109 |
before_set2pred_simp_tac i THEN_MAYBE |
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
|
110 |
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
|
111 |
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
|
112 |
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
|
113 |
dtac CollectD 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
|
114 |
TRY (split_all_tac i) THEN_MAYBE |
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
|
115 |
(rename_tac var_names i THEN full_simp_tac (HOL_basic_ss addsimps [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
|
116 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
117 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
118 |
(** 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
|
119 |
(** 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
|
120 |
(** 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
|
121 |
(** 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
|
122 |
(** 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
|
123 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
124 |
|
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
|
125 |
fun MaxSimpTac var_names tac = FIRST'[rtac subset_refl, set2pred_tac 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
|
126 |
|
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 |
fun BasicSimpTac var_names 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
|
128 |
simp_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
|
129 |
(HOL_basic_ss addsimps [mem_Collect_eq, split_conv] addsimprocs [record_simproc]) |
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 |
THEN_MAYBE' MaxSimpTac 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
|
131 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
132 |
|
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
|
133 |
(** 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
|
134 |
|
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 |
fun hoare_rule_tac (vars, Mlem) 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
|
136 |
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
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
((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
|
142 |
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
|
143 |
(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
|
144 |
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
|
145 |
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
|
146 |
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
|
147 |
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
|
148 |
rtac Mlem 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
|
149 |
split_simp_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
|
150 |
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
|
151 |
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
|
152 |
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
|
153 |
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
|
154 |
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
|
155 |
rtac @{thm WhileRule} 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
|
156 |
BasicSimpTac var_names tac (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
|
157 |
rule_tac true (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
|
158 |
THEN (if pre_cond then BasicSimpTac var_names tac i else rtac subset_refl 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 |
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
|
160 |
|
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 |
|
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 |
(** 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
|
163 |
(** 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
|
164 |
|
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 hoare_tac ctxt (tac: int -> tactic) = SUBGOAL (fn (goal, 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 |
SELECT_GOAL (hoare_rule_tac (Mset ctxt goal) tac true 1) 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
|
167 |