| author | nipkow | 
| Tue, 14 Apr 2015 16:47:55 +0200 | |
| changeset 60070 | 73c6e58a105c | 
| parent 58889 | 5b7a9633cfa8 | 
| child 63120 | 629a4c5e953e | 
| 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  | 
|
| 58889 | 6  | 
section{*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 = {*
 | 
|
| 55111 | 19  | 
Args.goal_spec -- Scan.lift Args.name_inner_syntax >>  | 
| 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 {*
 | 
| 
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}))
 | 
| 
42795
 
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  |