author | bulwahn |
Mon, 11 May 2009 09:18:42 +0200 | |
changeset 31106 | 9a1178204dc0 |
parent 30549 | d2d7874648bd |
child 32149 | ef59550a55d3 |
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 = {* |
30549 | 13 |
Scan.succeed (fn ctxt => |
30510
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents:
27882
diff
changeset
|
14 |
SIMPLE_METHOD' (constrains_tac (local_clasimpset_of ctxt))) *} |
13786 | 15 |
"for proving safety properties" |
16 |
||
17 |
method_setup ensures_tac = {* |
|
30549 | 18 |
Args.goal_spec -- Scan.lift Args.name_source >> |
19 |
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac (local_clasimpset_of ctxt) s)) |
|
20 |
*} "for proving progress properties" |
|
13786 | 21 |
|
22 |
end |