| author | nipkow | 
| Tue, 07 May 2013 10:35:40 +0200 | |
| changeset 51891 | b4e85748ce48 | 
| parent 51717 | 9e7d1c139569 | 
| child 58963 | 26bf09b95dda | 
| permissions | -rw-r--r-- | 
| 13786 | 1  | 
(* Title: HOL/UNITY/UNITY_tactics.ML  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 2003 University of Cambridge  | 
|
4  | 
||
| 24147 | 5  | 
Specialized UNITY tactics.  | 
| 13786 | 6  | 
*)  | 
7  | 
||
| 13797 | 8  | 
(*Combines two invariance ASSUMPTIONS into one. USEFUL??*)  | 
| 24147 | 9  | 
val Always_Int_tac = dtac @{thm Always_Int_I} THEN' assume_tac THEN' etac @{thm Always_thin}
 | 
| 13797 | 10  | 
|
11  | 
(*Combines a list of invariance THEOREMS into one.*)  | 
|
| 24147 | 12  | 
val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I})
 | 
| 13797 | 13  | 
|
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13798 
diff
changeset
 | 
14  | 
(*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
 | 
15  | 
(from weak constrains) must have been done already.*)  | 
| 42767 | 16  | 
fun constrains_tac ctxt i =  | 
| 42793 | 17  | 
SELECT_GOAL  | 
18  | 
(EVERY  | 
|
19  | 
[REPEAT (Always_Int_tac 1),  | 
|
20  | 
(*reduce the fancy safety properties to "constrains"*)  | 
|
21  | 
      REPEAT (etac @{thm Always_ConstrainsI} 1
 | 
|
22  | 
ORELSE  | 
|
23  | 
              resolve_tac [@{thm StableI}, @{thm stableI},
 | 
|
24  | 
                           @{thm constrains_imp_Constrains}] 1),
 | 
|
25  | 
(*for safety, the totalize operator can be ignored*)  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
26  | 
simp_tac (put_simpset HOL_ss ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
27  | 
        addsimps [@{thm mk_total_program_def}, @{thm totalize_constrains_iff}]) 1,
 | 
| 42793 | 28  | 
      rtac @{thm constrainsI} 1,
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
29  | 
full_simp_tac ctxt 1,  | 
| 42793 | 30  | 
REPEAT (FIRSTGOAL (etac disjE)),  | 
31  | 
ALLGOALS (clarify_tac ctxt),  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
32  | 
ALLGOALS (asm_full_simp_tac ctxt)]) i;  | 
| 13786 | 33  | 
|
34  | 
(*proves "ensures/leadsTo" properties when the program is specified*)  | 
|
| 42767 | 35  | 
fun ensures_tac ctxt sact =  | 
| 42793 | 36  | 
SELECT_GOAL  | 
37  | 
(EVERY  | 
|
38  | 
[REPEAT (Always_Int_tac 1),  | 
|
39  | 
      etac @{thm Always_LeadsTo_Basis} 1
 | 
|
40  | 
ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*)  | 
|
41  | 
          REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
 | 
|
42  | 
                            @{thm EnsuresI}, @{thm ensuresI}] 1),
 | 
|
43  | 
(*now there are two subgoals: co & transient*)  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
44  | 
      simp_tac (ctxt addsimps [@{thm mk_total_program_def}]) 2,
 | 
| 42793 | 45  | 
res_inst_tac ctxt  | 
46  | 
        [(("act", 0), sact)] @{thm totalize_transientI} 2
 | 
|
47  | 
ORELSE res_inst_tac ctxt  | 
|
48  | 
        [(("act", 0), sact)] @{thm transientI} 2,
 | 
|
49  | 
(*simplify the command's domain*)  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
50  | 
      simp_tac (ctxt addsimps @{thms Domain_unfold}) 3,
 | 
| 42793 | 51  | 
constrains_tac ctxt 1,  | 
52  | 
ALLGOALS (clarify_tac ctxt),  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
53  | 
ALLGOALS (asm_lr_simp_tac ctxt)]);  | 
| 13796 | 54  | 
|
| 13786 | 55  | 
|
56  | 
(*Composition equivalences, from Lift_prog*)  | 
|
57  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
58  | 
fun make_o_equivs ctxt th =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
59  | 
[th,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
60  | 
   th RS @{thm o_equiv_assoc} |> simplify (put_simpset HOL_ss ctxt addsimps [@{thm o_assoc}]),
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46752 
diff
changeset
 | 
61  | 
   th RS @{thm o_equiv_apply} |> simplify (put_simpset HOL_ss ctxt addsimps [@{thm o_def}, @{thm sub_def}])];
 | 
| 13786 | 62  |