author | wenzelm |
Wed, 28 Dec 2022 12:30:18 +0100 | |
changeset 76798 | 69d8d16c5612 |
parent 74563 | 042041c0ebeb |
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 |
|
63146 | 6 |
section\<open>Comprehensive UNITY Theory\<close> |
13786 | 7 |
|
32689 | 8 |
theory UNITY_Main |
9 |
imports Detects PPROD Follows ProgressSets |
|
10 |
begin |
|
13786 | 11 |
|
69605 | 12 |
ML_file \<open>UNITY_tactics.ML\<close> |
48891 | 13 |
|
63146 | 14 |
method_setup safety = \<open> |
15 |
Scan.succeed (SIMPLE_METHOD' o constrains_tac)\<close> |
|
13786 | 16 |
"for proving safety properties" |
17 |
||
63146 | 18 |
method_setup ensures_tac = \<open> |
74563 | 19 |
Args.goal_spec -- Scan.lift Parse.embedded_inner_syntax >> |
42767 | 20 |
(fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac ctxt s)) |
63146 | 21 |
\<close> "for proving progress properties" |
13786 | 22 |
|
63146 | 23 |
setup \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48891
diff
changeset
|
24 |
map_theory_simpset (fn ctxt => ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48891
diff
changeset
|
25 |
addsimps (make_o_equivs ctxt @{thm fst_o_funPair} @ make_o_equivs ctxt @{thm snd_o_funPair}) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48891
diff
changeset
|
26 |
addsimps (make_o_equivs ctxt @{thm fst_o_lift_map} @ make_o_equivs ctxt @{thm snd_o_lift_map})) |
63146 | 27 |
\<close> |
42795
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
wenzelm
parents:
42767
diff
changeset
|
28 |
|
13786 | 29 |
end |