| author | immler@in.tum.de | 
| Wed, 03 Jun 2009 16:56:41 +0200 | |
| changeset 31409 | d8537ba165b5 | 
| parent 30610 | bcbc34cb9749 | 
| child 34234 | 86a985e9b4df | 
| 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: 
13798diff
changeset | 15 | (*Proves "co" properties when the program is specified. Any use of invariants | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
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: 
13798diff
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: 
13798diff
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,
 | |
| 27239 | 44 | res_inst_tac (Simplifier.the_context ss) | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
24147diff
changeset | 45 |                 [(("act", 0), sact)] @{thm totalize_transientI} 2
 | 
| 27239 | 46 | ORELSE res_inst_tac (Simplifier.the_context ss) | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
24147diff
changeset | 47 |                 [(("act", 0), sact)] @{thm transientI} 2,
 | 
| 13786 | 48 | (*simplify the command's domain*) | 
| 24147 | 49 |               simp_tac (ss addsimps [@{thm Domain_def}]) 3,
 | 
| 50 | constrains_tac (cs,ss) 1, | |
| 51 | ALLGOALS (clarify_tac cs), | |
| 52 | ALLGOALS (asm_lr_simp_tac ss)]); | |
| 13796 | 53 | |
| 13786 | 54 | |
| 55 | (*Composition equivalences, from Lift_prog*) | |
| 56 | ||
| 57 | fun make_o_equivs th = | |
| 58 | [th, | |
| 24147 | 59 |      th RS @{thm o_equiv_assoc} |> simplify (HOL_ss addsimps [@{thm o_assoc}]),
 | 
| 60 |      th RS @{thm o_equiv_apply} |> simplify (HOL_ss addsimps [@{thm o_def}, @{thm sub_def}])];
 | |
| 13786 | 61 | |
| 30610 | 62 | Simplifier.change_simpset (fn ss => ss | 
| 63 |   addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair})
 | |
| 64 |   addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map}));
 |