author | haftmann |
Tue, 16 Oct 2007 23:12:45 +0200 | |
changeset 25062 | af5ef0d4d655 |
parent 24147 | edc90be09ac1 |
child 27882 | eaa9fef9f4c1 |
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:
15032
diff
changeset
|
12 |
method_setup safety = {* |
13786 | 13 |
Method.ctxt_args (fn ctxt => |
24147 | 14 |
Method.SIMPLE_METHOD' (constrains_tac (local_clasimpset_of ctxt))) *} |
13786 | 15 |
"for proving safety properties" |
16 |
||
17 |
method_setup ensures_tac = {* |
|
18 |
fn args => fn ctxt => |
|
21588 | 19 |
Method.goal_args' (Scan.lift Args.name) |
24147 | 20 |
(ensures_tac (local_clasimpset_of ctxt)) |
13786 | 21 |
args ctxt *} |
22 |
"for proving progress properties" |
|
23 |
||
24 |
end |