| author | urbanc |
| Fri, 04 Jan 2008 16:35:22 +0100 | |
| changeset 25832 | 41a014cc44c0 |
| parent 24147 | edc90be09ac1 |
| child 27208 | 5fe899199f85 |
| permissions | -rw-r--r-- |
| 13786 | 1 |
(* Title: HOL/UNITY/UNITY_tactics.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 2003 University of Cambridge |
|
5 |
||
| 24147 | 6 |
Specialized UNITY tactics. |
| 13786 | 7 |
*) |
8 |
||
| 13797 | 9 |
(*Combines two invariance ASSUMPTIONS into one. USEFUL??*) |
| 24147 | 10 |
val Always_Int_tac = dtac @{thm Always_Int_I} THEN' assume_tac THEN' etac @{thm Always_thin}
|
| 13797 | 11 |
|
12 |
(*Combines a list of invariance THEOREMS into one.*) |
|
| 24147 | 13 |
val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I})
|
| 13797 | 14 |
|
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13798
diff
changeset
|
15 |
(*Proves "co" properties when the program is specified. Any use of invariants |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13798
diff
changeset
|
16 |
(from weak constrains) must have been done already.*) |
| 24147 | 17 |
fun constrains_tac(cs,ss) i = |
| 13786 | 18 |
SELECT_GOAL |
19 |
(EVERY [REPEAT (Always_Int_tac 1), |
|
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13798
diff
changeset
|
20 |
(*reduce the fancy safety properties to "constrains"*) |
| 24147 | 21 |
REPEAT (etac @{thm Always_ConstrainsI} 1
|
22 |
ORELSE |
|
23 |
resolve_tac [@{thm StableI}, @{thm stableI},
|
|
24 |
@{thm constrains_imp_Constrains}] 1),
|
|
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13798
diff
changeset
|
25 |
(*for safety, the totalize operator can be ignored*) |
| 24147 | 26 |
simp_tac (HOL_ss addsimps |
27 |
[@{thm mk_total_program_def}, @{thm totalize_constrains_iff}]) 1,
|
|
28 |
rtac @{thm constrainsI} 1,
|
|
29 |
full_simp_tac ss 1, |
|
30 |
REPEAT (FIRSTGOAL (etac disjE)), |
|
31 |
ALLGOALS (clarify_tac cs), |
|
32 |
ALLGOALS (asm_lr_simp_tac ss)]) i; |
|
| 13786 | 33 |
|
34 |
(*proves "ensures/leadsTo" properties when the program is specified*) |
|
| 24147 | 35 |
fun ensures_tac(cs,ss) sact = |
| 13786 | 36 |
SELECT_GOAL |
37 |
(EVERY [REPEAT (Always_Int_tac 1), |
|
| 24147 | 38 |
etac @{thm Always_LeadsTo_Basis} 1
|
39 |
ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) |
|
40 |
REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
|
|
41 |
@{thm EnsuresI}, @{thm ensuresI}] 1),
|
|
42 |
(*now there are two subgoals: co & transient*) |
|
43 |
simp_tac (ss addsimps [@{thm mk_total_program_def}]) 2,
|
|
44 |
res_inst_tac [("act", sact)] @{thm totalize_transientI} 2
|
|
45 |
ORELSE res_inst_tac [("act", sact)] @{thm transientI} 2,
|
|
| 13786 | 46 |
(*simplify the command's domain*) |
| 24147 | 47 |
simp_tac (ss addsimps [@{thm Domain_def}]) 3,
|
48 |
constrains_tac (cs,ss) 1, |
|
49 |
ALLGOALS (clarify_tac cs), |
|
50 |
ALLGOALS (asm_lr_simp_tac ss)]); |
|
| 13796 | 51 |
|
| 13786 | 52 |
|
53 |
(*Composition equivalences, from Lift_prog*) |
|
54 |
||
55 |
fun make_o_equivs th = |
|
56 |
[th, |
|
| 24147 | 57 |
th RS @{thm o_equiv_assoc} |> simplify (HOL_ss addsimps [@{thm o_assoc}]),
|
58 |
th RS @{thm o_equiv_apply} |> simplify (HOL_ss addsimps [@{thm o_def}, @{thm sub_def}])];
|
|
| 13786 | 59 |
|
| 24147 | 60 |
Addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair});
|
| 13786 | 61 |
|
| 24147 | 62 |
Addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map});
|