author | wenzelm |
Wed, 22 Aug 2012 22:55:41 +0200 | |
changeset 48891 | c0eafbd55de3 |
parent 42795 | 66fcc9882784 |
child 51717 | 9e7d1c139569 |
permissions | -rw-r--r-- |
13786 | 1 |
(* Title: HOL/UNITY/UNITY_Main.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
Copyright 2003 University of Cambridge |
|
13798 | 4 |
*) |
13786 | 5 |
|
13798 | 6 |
header{*Comprehensive UNITY Theory*} |
13786 | 7 |
|
32689 | 8 |
theory UNITY_Main |
9 |
imports Detects PPROD Follows ProgressSets |
|
10 |
begin |
|
13786 | 11 |
|
48891 | 12 |
ML_file "UNITY_tactics.ML" |
13 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
15032
diff
changeset
|
14 |
method_setup safety = {* |
42767 | 15 |
Scan.succeed (SIMPLE_METHOD' o constrains_tac) *} |
13786 | 16 |
"for proving safety properties" |
17 |
||
18 |
method_setup ensures_tac = {* |
|
30549 | 19 |
Args.goal_spec -- Scan.lift Args.name_source >> |
42767 | 20 |
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac ctxt s)) |
30549 | 21 |
*} "for proving progress properties" |
13786 | 22 |
|
42795
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
23 |
setup {* |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
24 |
Simplifier.map_simpset_global (fn ss => ss |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
25 |
addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair}) |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
26 |
addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map})) |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
27 |
*} |
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
28 |
|
13786 | 29 |
end |