| author | wenzelm |
| Fri, 28 Aug 2009 18:23:24 +0200 | |
| changeset 32431 | bcd14373ec30 |
| parent 32149 | ef59550a55d3 |
| child 32689 | 860e1a2317bd |
| 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 => |
|
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
30549
diff
changeset
|
14 |
SIMPLE_METHOD' (constrains_tac (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 >> |
|
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
30549
diff
changeset
|
19 |
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac (clasimpset_of ctxt) s)) |
| 30549 | 20 |
*} "for proving progress properties" |
| 13786 | 21 |
|
22 |
end |