| author | haftmann | 
| Sun, 17 Jul 2011 20:46:51 +0200 | |
| changeset 43871 | 79c3231e0593 | 
| parent 42795 | 66fcc9882784 | 
| child 48891 | c0eafbd55de3 | 
| permissions | -rw-r--r-- | 
| 13786 | 1 | (* Title: HOL/UNITY/UNITY_Main.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 2003 University of Cambridge | |
| 13798 | 4 | *) | 
| 13786 | 5 | |
| 13798 | 6 | header{*Comprehensive UNITY Theory*}
 | 
| 13786 | 7 | |
| 32689 | 8 | theory UNITY_Main | 
| 9 | imports Detects PPROD Follows ProgressSets | |
| 10 | uses "UNITY_tactics.ML" | |
| 11 | begin | |
| 13786 | 12 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
15032diff
changeset | 13 | method_setup safety = {*
 | 
| 42767 | 14 | Scan.succeed (SIMPLE_METHOD' o constrains_tac) *} | 
| 13786 | 15 | "for proving safety properties" | 
| 16 | ||
| 17 | method_setup ensures_tac = {*
 | |
| 30549 | 18 | Args.goal_spec -- Scan.lift Args.name_source >> | 
| 42767 | 19 | (fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac ctxt s)) | 
| 30549 | 20 | *} "for proving progress properties" | 
| 13786 | 21 | |
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 22 | setup {*
 | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 23 | Simplifier.map_simpset_global (fn ss => ss | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 24 |     addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair})
 | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 25 |     addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map}))
 | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 26 | *} | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42767diff
changeset | 27 | |
| 13786 | 28 | end |