| author | wenzelm | 
| Tue, 16 May 2006 13:01:23 +0200 | |
| changeset 19641 | f1de44e61ec1 | 
| parent 16417 | 9bc16273c2d4 | 
| child 21588 | cd0dc678a205 | 
| permissions | -rw-r--r-- | 
| 13786 | 1 | (* Title: HOL/UNITY/UNITY_Main.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 2003 University of Cambridge | |
| 13798 | 5 | *) | 
| 13786 | 6 | |
| 13798 | 7 | header{*Comprehensive UNITY Theory*}
 | 
| 13786 | 8 | |
| 16417 | 9 | theory UNITY_Main imports Detects PPROD Follows ProgressSets | 
| 10 | uses "UNITY_tactics.ML" begin | |
| 13786 | 11 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
15032diff
changeset | 12 | method_setup safety = {*
 | 
| 13786 | 13 | Method.ctxt_args (fn ctxt => | 
| 14 | Method.METHOD (fn facts => | |
| 15032 | 15 | gen_constrains_tac (local_clasimpset_of ctxt) 1)) *} | 
| 13786 | 16 | "for proving safety properties" | 
| 17 | ||
| 18 | method_setup ensures_tac = {*
 | |
| 19 | fn args => fn ctxt => | |
| 20 | Method.goal_args' (Scan.lift Args.name) | |
| 15032 | 21 | (gen_ensures_tac (local_clasimpset_of ctxt)) | 
| 13786 | 22 | args ctxt *} | 
| 23 | "for proving progress properties" | |
| 24 | ||
| 25 | end |