| author | wenzelm | 
| Fri, 17 Aug 2007 23:10:45 +0200 | |
| changeset 24313 | 5a6342236a32 | 
| parent 22095 | 07875394618e | 
| child 26425 | 6561665c5cb1 | 
| permissions | -rw-r--r-- | 
| 9772 | 1  | 
(* Title: Provers/clasimp.ML  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
3  | 
Author: David von Oheimb, TU Muenchen  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
4  | 
|
| 5219 | 5  | 
Combination of classical reasoner and simplifier (depends on  | 
| 16019 | 6  | 
splitter.ML, classical.ML, blast.ML).  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
8  | 
|
| 9860 | 9  | 
infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 addsimps2 delsimps2  | 
| 
13603
 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 
berghofe 
parents: 
13026 
diff
changeset
 | 
10  | 
addSss addss addss' addIffs delIffs;  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
11  | 
|
| 5219 | 12  | 
signature CLASIMP_DATA =  | 
13  | 
sig  | 
|
| 8469 | 14  | 
structure Splitter: SPLITTER  | 
| 5219 | 15  | 
structure Classical: CLASSICAL  | 
16  | 
structure Blast: BLAST  | 
|
17  | 
sharing type Classical.claset = Blast.claset  | 
|
| 9860 | 18  | 
val notE: thm  | 
19  | 
val iffD1: thm  | 
|
20  | 
val iffD2: thm  | 
|
| 5219 | 21  | 
end  | 
22  | 
||
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
23  | 
signature CLASIMP =  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
24  | 
sig  | 
| 5219 | 25  | 
include CLASIMP_DATA  | 
26  | 
type claset  | 
|
27  | 
type clasimpset  | 
|
| 17879 | 28  | 
val change_clasimpset: (clasimpset -> clasimpset) -> unit  | 
| 10317 | 29  | 
val clasimpset: unit -> clasimpset  | 
| 9860 | 30  | 
val addSIs2: clasimpset * thm list -> clasimpset  | 
31  | 
val addSEs2: clasimpset * thm list -> clasimpset  | 
|
32  | 
val addSDs2: clasimpset * thm list -> clasimpset  | 
|
33  | 
val addIs2: clasimpset * thm list -> clasimpset  | 
|
34  | 
val addEs2: clasimpset * thm list -> clasimpset  | 
|
35  | 
val addDs2: clasimpset * thm list -> clasimpset  | 
|
36  | 
val addsimps2: clasimpset * thm list -> clasimpset  | 
|
37  | 
val delsimps2: clasimpset * thm list -> clasimpset  | 
|
38  | 
val addSss: claset * simpset -> claset  | 
|
39  | 
val addss: claset * simpset -> claset  | 
|
| 
13603
 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 
berghofe 
parents: 
13026 
diff
changeset
 | 
40  | 
val addss': claset * simpset -> claset  | 
| 9860 | 41  | 
val addIffs: clasimpset * thm list -> clasimpset  | 
42  | 
val delIffs: clasimpset * thm list -> clasimpset  | 
|
43  | 
val AddIffs: thm list -> unit  | 
|
44  | 
val DelIffs: thm list -> unit  | 
|
45  | 
val CLASIMPSET: (clasimpset -> tactic) -> tactic  | 
|
46  | 
val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic  | 
|
| 5483 | 47  | 
val clarsimp_tac: clasimpset -> int -> tactic  | 
| 9860 | 48  | 
val Clarsimp_tac: int -> tactic  | 
49  | 
val mk_auto_tac: clasimpset -> int -> int -> tactic  | 
|
50  | 
val auto_tac: clasimpset -> tactic  | 
|
51  | 
val Auto_tac: tactic  | 
|
52  | 
val auto: unit -> unit  | 
|
53  | 
val force_tac: clasimpset -> int -> tactic  | 
|
54  | 
val Force_tac: int -> tactic  | 
|
55  | 
val force: int -> unit  | 
|
56  | 
val fast_simp_tac: clasimpset -> int -> tactic  | 
|
57  | 
val slow_simp_tac: clasimpset -> int -> tactic  | 
|
58  | 
val best_simp_tac: clasimpset -> int -> tactic  | 
|
| 18728 | 59  | 
val attrib: (clasimpset * thm list -> clasimpset) -> attribute  | 
| 9506 | 60  | 
val get_local_clasimpset: Proof.context -> clasimpset  | 
| 15032 | 61  | 
val local_clasimpset_of: Proof.context -> clasimpset  | 
| 18728 | 62  | 
val iff_add: attribute  | 
63  | 
val iff_add': attribute  | 
|
64  | 
val iff_del: attribute  | 
|
| 9860 | 65  | 
val iff_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list  | 
| 9506 | 66  | 
val clasimp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list  | 
| 18708 | 67  | 
val setup: theory -> theory  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
68  | 
end;  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
69  | 
|
| 5219 | 70  | 
functor ClasimpFun(Data: CLASIMP_DATA): CLASIMP =  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
71  | 
struct  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
72  | 
|
| 5219 | 73  | 
open Data;  | 
74  | 
||
75  | 
type claset = Classical.claset;  | 
|
76  | 
type clasimpset = claset * simpset;  | 
|
77  | 
||
| 17879 | 78  | 
fun change_clasimpset_of thy f =  | 
79  | 
let val (cs', ss') = f (Classical.get_claset thy, Simplifier.get_simpset thy) in  | 
|
80  | 
Classical.change_claset_of thy (fn _ => cs');  | 
|
81  | 
Simplifier.change_simpset_of thy (fn _ => ss')  | 
|
82  | 
end;  | 
|
83  | 
||
| 
22095
 
07875394618e
moved ML context stuff to from Context to ML_Context;
 
wenzelm 
parents: 
21963 
diff
changeset
 | 
84  | 
fun change_clasimpset f = change_clasimpset_of (ML_Context.the_context ()) f;  | 
| 10317 | 85  | 
fun clasimpset () = (Classical.claset (), Simplifier.simpset ());  | 
86  | 
||
| 5219 | 87  | 
|
88  | 
(* clasimpset operations *)  | 
|
89  | 
||
90  | 
(*this interface for updating a clasimpset is rudimentary and just for  | 
|
91  | 
convenience for the most common cases*)  | 
|
92  | 
||
93  | 
fun pair_upd1 f ((a,b),x) = (f(a,x), b);  | 
|
94  | 
fun pair_upd2 f ((a,b),x) = (a, f(b,x));  | 
|
95  | 
||
96  | 
fun op addSIs2 arg = pair_upd1 Classical.addSIs arg;  | 
|
97  | 
fun op addSEs2 arg = pair_upd1 Classical.addSEs arg;  | 
|
98  | 
fun op addSDs2 arg = pair_upd1 Classical.addSDs arg;  | 
|
99  | 
fun op addIs2 arg = pair_upd1 Classical.addIs arg;  | 
|
100  | 
fun op addEs2 arg = pair_upd1 Classical.addEs arg;  | 
|
101  | 
fun op addDs2 arg = pair_upd1 Classical.addDs arg;  | 
|
102  | 
fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg;  | 
|
103  | 
fun op delsimps2 arg = pair_upd2 Simplifier.delsimps arg;  | 
|
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
104  | 
|
| 9402 | 105  | 
(*not totally safe: may instantiate unknowns that appear also in other subgoals*)  | 
106  | 
val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true);  | 
|
107  | 
||
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
108  | 
(*Add a simpset to a classical set!*)  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
109  | 
(*Caution: only one simpset added can be added by each of addSss and addss*)  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
10821 
diff
changeset
 | 
110  | 
fun cs addSss ss = Classical.addSafter (cs, ("safe_asm_full_simp_tac",
 | 
| 9772 | 111  | 
CHANGED o safe_asm_full_simp_tac ss));  | 
112  | 
fun cs addss  ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
 | 
|
113  | 
CHANGED o Simplifier.asm_full_simp_tac ss));  | 
|
| 
13603
 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 
berghofe 
parents: 
13026 
diff
changeset
 | 
114  | 
fun cs addss' ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
 | 
| 
 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 
berghofe 
parents: 
13026 
diff
changeset
 | 
115  | 
CHANGED o Simplifier.asm_lr_simp_tac ss));  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
116  | 
|
| 9860 | 117  | 
(* iffs: addition of rules to simpsets and clasets simultaneously *)  | 
118  | 
||
| 11344 | 119  | 
(*Takes (possibly conditional) theorems of the form A<->B to  | 
| 9860 | 120  | 
the Safe Intr rule B==>A and  | 
121  | 
the Safe Destruct rule A==>B.  | 
|
122  | 
Also ~A goes to the Safe Elim rule A ==> ?R  | 
|
| 11462 | 123  | 
Failing other cases, A is added as a Safe Intr rule*)  | 
| 9860 | 124  | 
local  | 
125  | 
||
| 13026 | 126  | 
fun addIff decls1 decls2 simp ((cs, ss), th) =  | 
127  | 
let  | 
|
128  | 
val n = nprems_of th;  | 
|
| 
17084
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
16019 
diff
changeset
 | 
129  | 
val (elim, intro) = if n = 0 then decls1 else decls2;  | 
| 13026 | 130  | 
val zero_rotate = zero_var_indexes o rotate_prems n;  | 
131  | 
in  | 
|
| 
21709
 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 
paulson 
parents: 
21646 
diff
changeset
 | 
132  | 
(elim (intro (cs, [zero_rotate (th RS iffD2)]),  | 
| 
 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 
paulson 
parents: 
21646 
diff
changeset
 | 
133  | 
[Tactic.make_elim (zero_rotate (th RS iffD1))])  | 
| 
 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 
paulson 
parents: 
21646 
diff
changeset
 | 
134  | 
handle THM _ => (elim (cs, [zero_rotate (th RS notE)])  | 
| 11902 | 135  | 
handle THM _ => intro (cs, [th])),  | 
136  | 
simp (ss, [th]))  | 
|
137  | 
end;  | 
|
| 9860 | 138  | 
|
| 12375 | 139  | 
fun delIff delcs delss ((cs, ss), th) =  | 
| 11902 | 140  | 
let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in  | 
| 18688 | 141  | 
(delcs (cs, [zero_rotate (th RS iffD2),  | 
142  | 
Tactic.make_elim (zero_rotate (th RS iffD1))])  | 
|
143  | 
handle THM _ => (delcs (cs, [zero_rotate (th RS notE)])  | 
|
| 12375 | 144  | 
handle THM _ => delcs (cs, [th])),  | 
145  | 
delss (ss, [th]))  | 
|
| 11902 | 146  | 
end;  | 
| 9860 | 147  | 
|
| 18630 | 148  | 
fun modifier att (x, ths) =  | 
| 18728 | 149  | 
fst (foldl_map (Library.apply [att]) (x, rev ths));  | 
| 18630 | 150  | 
|
| 18688 | 151  | 
val addXIs = modifier (ContextRules.intro_query NONE);  | 
152  | 
val addXEs = modifier (ContextRules.elim_query NONE);  | 
|
153  | 
val addXDs = modifier (ContextRules.dest_query NONE);  | 
|
154  | 
val delXs = modifier ContextRules.rule_del;  | 
|
| 18630 | 155  | 
|
| 9860 | 156  | 
in  | 
157  | 
||
| 
10033
 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 
wenzelm 
parents: 
9952 
diff
changeset
 | 
158  | 
val op addIffs =  | 
| 
17084
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
16019 
diff
changeset
 | 
159  | 
Library.foldl  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
16019 
diff
changeset
 | 
160  | 
(addIff (Classical.addSEs, Classical.addSIs)  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
16019 
diff
changeset
 | 
161  | 
(Classical.addEs, Classical.addIs)  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
16019 
diff
changeset
 | 
162  | 
Simplifier.addsimps);  | 
| 15570 | 163  | 
val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps);  | 
| 
10033
 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 
wenzelm 
parents: 
9952 
diff
changeset
 | 
164  | 
|
| 17879 | 165  | 
fun AddIffs thms = change_clasimpset (fn css => css addIffs thms);  | 
166  | 
fun DelIffs thms = change_clasimpset (fn css => css delIffs thms);  | 
|
| 9860 | 167  | 
|
| 18688 | 168  | 
fun addIffs_generic (x, ths) =  | 
169  | 
Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1;  | 
|
| 12375 | 170  | 
|
| 18688 | 171  | 
fun delIffs_generic (x, ths) =  | 
172  | 
Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1;  | 
|
| 12375 | 173  | 
|
| 9860 | 174  | 
end;  | 
175  | 
||
176  | 
||
| 5219 | 177  | 
(* tacticals *)  | 
178  | 
||
179  | 
fun CLASIMPSET tacf state =  | 
|
180  | 
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state;  | 
|
181  | 
||
182  | 
fun CLASIMPSET' tacf i state =  | 
|
183  | 
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state;  | 
|
| 4888 | 184  | 
|
185  | 
||
| 12375 | 186  | 
fun clarsimp_tac (cs, ss) =  | 
187  | 
safe_asm_full_simp_tac ss THEN_ALL_NEW  | 
|
188  | 
Classical.clarify_tac (cs addSss ss);  | 
|
189  | 
||
| 10317 | 190  | 
fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i;  | 
| 5483 | 191  | 
|
192  | 
||
| 5219 | 193  | 
(* auto_tac *)  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
194  | 
|
| 5219 | 195  | 
fun blast_depth_tac cs m i thm =  | 
| 9772 | 196  | 
Blast.depth_tac cs m i thm  | 
| 5554 | 197  | 
      handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
 | 
| 9772 | 198  | 
|
199  | 
(* a variant of depth_tac that avoids interference of the simplifier  | 
|
| 5219 | 200  | 
with dup_step_tac when they are combined by auto_tac *)  | 
| 
5756
 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5567 
diff
changeset
 | 
201  | 
local  | 
| 9772 | 202  | 
fun slow_step_tac' cs = Classical.appWrappers cs  | 
203  | 
(Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs);  | 
|
204  | 
in fun nodup_depth_tac cs m i state = SELECT_GOAL  | 
|
205  | 
(Classical.safe_steps_tac cs 1 THEN_ELSE  | 
|
206  | 
(DEPTH_SOLVE (nodup_depth_tac cs m 1),  | 
|
207  | 
Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac  | 
|
208  | 
(slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m-1) 1))  | 
|
| 
5756
 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5567 
diff
changeset
 | 
209  | 
)) i state;  | 
| 
 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5567 
diff
changeset
 | 
210  | 
end;  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
211  | 
|
| 9402 | 212  | 
(*Designed to be idempotent, except if blast_depth_tac instantiates variables  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
213  | 
in some of the subgoals*)  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
214  | 
fun mk_auto_tac (cs, ss) m n =  | 
| 5219 | 215  | 
let val cs' = cs addss ss  | 
| 9772 | 216  | 
val maintac =  | 
217  | 
blast_depth_tac cs m (* fast but can't use wrappers *)  | 
|
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
218  | 
ORELSE'  | 
| 
5756
 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5567 
diff
changeset
 | 
219  | 
(CHANGED o nodup_depth_tac cs' n); (* slower but more general *)  | 
| 5219 | 220  | 
in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss),  | 
| 9772 | 221  | 
TRY (Classical.safe_tac cs),  | 
222  | 
REPEAT (FIRSTGOAL maintac),  | 
|
| 5219 | 223  | 
TRY (Classical.safe_tac (cs addSss ss)),  | 
| 9772 | 224  | 
prune_params_tac]  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
225  | 
end;  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
226  | 
|
| 9772 | 227  | 
fun auto_tac css = mk_auto_tac css 4 2;  | 
| 10317 | 228  | 
fun Auto_tac st = auto_tac (clasimpset ()) st;  | 
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
229  | 
fun auto () = by Auto_tac;  | 
| 
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
230  | 
|
| 9772 | 231  | 
|
| 5219 | 232  | 
(* force_tac *)  | 
233  | 
||
| 4659 | 234  | 
(* aimed to solve the given subgoal totally, using whatever tools possible *)  | 
| 
4717
 
1370ad043564
renamed smart_tac to force_tac, slight improvement of force_tac
 
oheimb 
parents: 
4659 
diff
changeset
 | 
235  | 
fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [  | 
| 9772 | 236  | 
Classical.clarify_tac cs' 1,  | 
237  | 
IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1),  | 
|
238  | 
ALLGOALS (Classical.first_best_tac cs')]) end;  | 
|
| 10317 | 239  | 
fun Force_tac i = force_tac (clasimpset ()) i;  | 
| 
4717
 
1370ad043564
renamed smart_tac to force_tac, slight improvement of force_tac
 
oheimb 
parents: 
4659 
diff
changeset
 | 
240  | 
fun force i = by (Force_tac i);  | 
| 4659 | 241  | 
|
| 5219 | 242  | 
|
| 9805 | 243  | 
(* basic combinations *)  | 
244  | 
||
245  | 
fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end;  | 
|
| 9591 | 246  | 
|
| 9805 | 247  | 
val fast_simp_tac = ADDSS Classical.fast_tac;  | 
248  | 
val slow_simp_tac = ADDSS Classical.slow_tac;  | 
|
249  | 
val best_simp_tac = ADDSS Classical.best_tac;  | 
|
| 9591 | 250  | 
|
251  | 
||
| 9860 | 252  | 
(* access clasimpset *)  | 
| 8639 | 253  | 
|
| 7153 | 254  | 
fun get_local_clasimpset ctxt =  | 
255  | 
(Classical.get_local_claset ctxt, Simplifier.get_local_simpset ctxt);  | 
|
| 5926 | 256  | 
|
| 15032 | 257  | 
fun local_clasimpset_of ctxt =  | 
258  | 
(Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt);  | 
|
259  | 
||
| 9860 | 260  | 
|
261  | 
(* attributes *)  | 
|
262  | 
||
| 18728 | 263  | 
fun attrib f = Thm.declaration_attribute (fn th =>  | 
| 18688 | 264  | 
fn Context.Theory thy => (change_clasimpset_of thy (fn css => f (css, [th])); Context.Theory thy)  | 
265  | 
| Context.Proof ctxt =>  | 
|
266  | 
let  | 
|
267  | 
val cs = Classical.get_local_claset ctxt;  | 
|
268  | 
val ss = Simplifier.get_local_simpset ctxt;  | 
|
269  | 
val (cs', ss') = f ((cs, ss), [th]);  | 
|
270  | 
val ctxt' =  | 
|
271  | 
ctxt  | 
|
272  | 
|> Classical.put_local_claset cs'  | 
|
273  | 
|> Simplifier.put_local_simpset ss';  | 
|
274  | 
in Context.Proof ctxt' end);  | 
|
| 9860 | 275  | 
|
| 18688 | 276  | 
fun attrib' f (x, th) = (f (x, [th]), th);  | 
| 
10033
 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 
wenzelm 
parents: 
9952 
diff
changeset
 | 
277  | 
|
| 18688 | 278  | 
val iff_add = attrib (op addIffs);  | 
279  | 
val iff_add' = attrib' addIffs_generic;  | 
|
280  | 
val iff_del = attrib (op delIffs) o attrib' delIffs_generic;  | 
|
281  | 
||
282  | 
val iff_att = Attrib.syntax (Scan.lift  | 
|
283  | 
(Args.del >> K iff_del ||  | 
|
284  | 
Scan.option Args.add -- Args.query >> K iff_add' ||  | 
|
285  | 
Scan.option Args.add >> K iff_add));  | 
|
| 9860 | 286  | 
|
287  | 
||
288  | 
(* method modifiers *)  | 
|
289  | 
||
290  | 
val iffN = "iff";  | 
|
291  | 
||
292  | 
val iff_modifiers =  | 
|
| 18728 | 293  | 
[Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier),  | 
294  | 
Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'),  | 
|
295  | 
Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)];  | 
|
| 9860 | 296  | 
|
| 8469 | 297  | 
val clasimp_modifiers =  | 
| 9860 | 298  | 
Simplifier.simp_modifiers @ Splitter.split_modifiers @  | 
299  | 
Classical.cla_modifiers @ iff_modifiers;  | 
|
300  | 
||
301  | 
||
302  | 
(* methods *)  | 
|
| 5926 | 303  | 
|
| 7559 | 304  | 
fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts =>  | 
| 15032 | 305  | 
ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt));  | 
| 7132 | 306  | 
|
| 7559 | 307  | 
fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts =>  | 
| 15032 | 308  | 
HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt)));  | 
| 5926 | 309  | 
|
| 7559 | 310  | 
val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth;  | 
311  | 
val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth';  | 
|
| 5926 | 312  | 
|
313  | 
||
| 9772 | 314  | 
fun auto_args m =  | 
315  | 
Method.bang_sectioned_args' clasimp_modifiers (Scan.lift (Scan.option (Args.nat -- Args.nat))) m;  | 
|
316  | 
||
| 15531 | 317  | 
fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac)  | 
318  | 
| auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n));  | 
|
| 9772 | 319  | 
|
320  | 
||
321  | 
(* theory setup *)  | 
|
322  | 
||
| 5926 | 323  | 
val setup =  | 
| 18708 | 324  | 
(Attrib.add_attributes  | 
| 18728 | 325  | 
   [("iff", iff_att, "declaration of Simplifier / Classical rules")] #>
 | 
| 9860 | 326  | 
Method.add_methods  | 
| 9591 | 327  | 
   [("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"),
 | 
| 9805 | 328  | 
    ("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"),
 | 
329  | 
    ("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"),
 | 
|
| 9591 | 330  | 
    ("force", clasimp_method' force_tac, "force"),
 | 
| 9772 | 331  | 
    ("auto", auto_args auto_meth, "auto"),
 | 
| 18708 | 332  | 
    ("clarsimp", clasimp_method' (CHANGED_PROP oo clarsimp_tac), "clarify simplified goal")]);
 | 
| 5926 | 333  | 
|
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents:  
diff
changeset
 | 
334  | 
end;  |