author | haftmann |
Tue, 16 Oct 2007 23:12:45 +0200 | |
changeset 25062 | af5ef0d4d655 |
parent 24475 | a297ae4ff188 |
child 26300 | 03def556e26e |
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 |
Copyright 1998 TUM |
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 |
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
|
7 |
*) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
8 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
9 |
(*** The tactics ***) |
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 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
12 |
(** 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
|
13 |
(** "?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
|
14 |
(** 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
|
15 |
(** 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
|
16 |
(*****************************************************************************) |
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 |
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
|
19 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
20 |
(** 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
|
21 |
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
|
22 |
| 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
|
23 |
| abs2list _ = []; |
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 |
(** 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
|
26 |
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
|
27 |
| mk_vars _ = []; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
28 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
29 |
(** abstraction of body over a tuple formed from a list of free variables. |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
30 |
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
|
31 |
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
|
32 |
| 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
|
33 |
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
|
34 |
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
|
35 |
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
|
36 |
| Const (_, Type (_,[_, Type (_,[T,_])])) $ _ => T; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
37 |
in Const ("split", (T --> T2 --> boolT) --> mk_prodT (T,T2) --> boolT) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
38 |
$ 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
|
39 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
40 |
(** 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
|
41 |
fun mk_bodyC [] = HOLogic.unit |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
42 |
| mk_bodyC (x::xs) = if xs=[] then x |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
43 |
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
|
44 |
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
|
45 |
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
|
46 |
| 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
|
47 |
("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
|
48 |
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
|
49 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
50 |
(** maps a goal of the form: |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
51 |
1. [| P |] ==> VARS x1 ... xn {._.} _ {._.} or 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
|
52 |
fun get_vars thm = let val c = Logic.unprotect (concl_of (thm)); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
53 |
val d = Logic.strip_assums_concl c; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
54 |
val Const _ $ pre $ _ $ _ = dest_Trueprop d; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
55 |
in mk_vars pre end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
56 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
57 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
58 |
(** Makes Collect with type **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
59 |
fun mk_CollectC trm = let val T as Type ("fun",[t,_]) = fastype_of trm |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
60 |
in Collect_const t $ trm end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
61 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
62 |
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
|
63 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
64 |
(** Makes "Mset <= t" **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
65 |
fun Mset_incl t = let val MsetT = fastype_of t |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
66 |
in mk_Trueprop ((inclt MsetT) $ Free ("Mset", MsetT) $ t) end; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
67 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
68 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
69 |
fun Mset thm = let val vars = get_vars(thm); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
70 |
val varsT = fastype_of (mk_bodyC vars); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
71 |
val big_Collect = mk_CollectC (mk_abstupleC vars |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
72 |
(Free ("P",varsT --> boolT) $ mk_bodyC vars)); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
73 |
val small_Collect = mk_CollectC (Abs("x",varsT, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
74 |
Free ("P",varsT --> boolT) $ Bound 0)); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
75 |
val impl = implies $ (Mset_incl big_Collect) $ |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
76 |
(Mset_incl small_Collect); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
77 |
in Goal.prove (ProofContext.init (Thm.theory_of_thm thm)) ["Mset", "P"] [] impl (K (CLASET' blast_tac 1)) end; |
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 |
Goal "-(Collect b) = {x. ~(b x)}"; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
90 |
by (Fast_tac 1); |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
91 |
qed "Compl_Collect"; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
92 |
|
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 |
(**Simp_tacs**) |
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 |
val before_set2pred_simp_tac = |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
97 |
(simp_tac (HOL_basic_ss addsimps [Collect_conj_eq RS sym,Compl_Collect])); |
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 |
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
|
100 |
|
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 |
(** set2pred transforms sets inclusion into predicates implication, **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
103 |
(** 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
|
104 |
(** 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
|
105 |
(** 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
|
106 |
(** 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
|
107 |
(** 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
|
108 |
(** transformed. **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
109 |
(** 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
|
110 |
(** 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
|
111 |
(*****************************************************************************) |
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 |
fun set2pred i thm = |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
114 |
let val var_names = map (fst o dest_Free) (get_vars thm) in |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
115 |
((before_set2pred_simp_tac i) THEN_MAYBE |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
116 |
(EVERY [rtac subsetI i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
117 |
rtac CollectI i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
118 |
dtac CollectD i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
119 |
(TRY(split_all_tac i)) THEN_MAYBE |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
120 |
((rename_params_tac var_names i) THEN |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
121 |
(full_simp_tac (HOL_basic_ss addsimps [split_conv]) i)) ])) thm |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
122 |
end; |
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 |
(*****************************************************************************) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
125 |
(** 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
|
126 |
(** 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
|
127 |
(** 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
|
128 |
(** 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
|
129 |
(** 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
|
130 |
(*****************************************************************************) |
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 |
fun MaxSimpTac tac = FIRST'[rtac subset_refl, set2pred THEN_MAYBE' tac]; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
133 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
134 |
fun BasicSimpTac tac = |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
135 |
simp_tac |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
136 |
(HOL_basic_ss addsimps [mem_Collect_eq,split_conv] addsimprocs [record_simproc]) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
137 |
THEN_MAYBE' MaxSimpTac tac; |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
138 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
139 |
(** HoareRuleTac **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
140 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
141 |
fun WlpTac Mlem tac i = |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
142 |
rtac @{thm SeqRule} i THEN HoareRuleTac Mlem tac false (i+1) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
143 |
and HoareRuleTac Mlem tac pre_cond i st = st |> |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
144 |
(*abstraction over st prevents looping*) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
145 |
( (WlpTac Mlem tac i THEN HoareRuleTac Mlem tac pre_cond i) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
146 |
ORELSE |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
147 |
(FIRST[rtac @{thm SkipRule} i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
148 |
EVERY[rtac @{thm BasicRule} i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
149 |
rtac Mlem i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
150 |
split_simp_tac i], |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
151 |
EVERY[rtac @{thm CondRule} i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
152 |
HoareRuleTac Mlem tac false (i+2), |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
153 |
HoareRuleTac Mlem tac false (i+1)], |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
154 |
EVERY[rtac @{thm WhileRule} i, |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
155 |
BasicSimpTac tac (i+2), |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
156 |
HoareRuleTac Mlem tac true (i+1)] ] |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
157 |
THEN (if pre_cond then (BasicSimpTac tac i) else (rtac subset_refl i)) )); |
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 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
160 |
(** tac:(int -> tactic) is the tactic the user chooses to solve or simplify **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
161 |
(** the final verification conditions **) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
162 |
|
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
163 |
fun hoare_tac tac i thm = |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
164 |
let val Mlem = Mset(thm) |
a297ae4ff188
added Hoare/hoare_tac.ML (code from Hoare/Hoare.thy, also required in Isar_examples/Hoare.thy);
wenzelm
parents:
diff
changeset
|
165 |
in SELECT_GOAL(EVERY[HoareRuleTac Mlem tac true 1]) i thm end; |