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
|
|
5 |
|
|
6 |
Inclusive UNITY theory
|
|
7 |
*)
|
|
8 |
|
|
9 |
theory UNITY_Main = Detects + PPROD + Follows
|
|
10 |
files "UNITY_tactics.ML":
|
|
11 |
|
|
12 |
method_setup constrains = {*
|
|
13 |
Method.ctxt_args (fn ctxt =>
|
|
14 |
Method.METHOD (fn facts =>
|
|
15 |
gen_constrains_tac (Classical.get_local_claset ctxt,
|
|
16 |
Simplifier.get_local_simpset ctxt) 1)) *}
|
|
17 |
"for proving safety properties"
|
|
18 |
|
|
19 |
method_setup ensures_tac = {*
|
|
20 |
fn args => fn ctxt =>
|
|
21 |
Method.goal_args' (Scan.lift Args.name)
|
|
22 |
(gen_ensures_tac (Classical.get_local_claset ctxt,
|
|
23 |
Simplifier.get_local_simpset ctxt))
|
|
24 |
args ctxt *}
|
|
25 |
"for proving progress properties"
|
|
26 |
|
|
27 |
end
|