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