author | wenzelm |
Thu, 07 Sep 2000 20:51:07 +0200 | |
changeset 9893 | 93d2fde0306c |
parent 9860 | 5c5efed691b9 |
child 9952 | 24914e42b857 |
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 |
8469 | 6 |
simplifier.ML, 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 |
10 |
addSss 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 |
|
14 |
structure Simplifier: SIMPLIFIER |
|
8469 | 15 |
structure Splitter: SPLITTER |
5219 | 16 |
structure Classical: CLASSICAL |
17 |
structure Blast: BLAST |
|
18 |
sharing type Classical.claset = Blast.claset |
|
9860 | 19 |
val dest_Trueprop: term -> term |
20 |
val not_const: term |
|
21 |
val iff_const: term |
|
22 |
val notE: thm |
|
23 |
val iffD1: thm |
|
24 |
val iffD2: thm |
|
25 |
val cla_make_elim: thm -> thm |
|
5219 | 26 |
end |
27 |
||
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
28 |
signature CLASIMP = |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
29 |
sig |
5219 | 30 |
include CLASIMP_DATA |
31 |
type claset |
|
32 |
type simpset |
|
33 |
type clasimpset |
|
9860 | 34 |
val addSIs2: clasimpset * thm list -> clasimpset |
35 |
val addSEs2: clasimpset * thm list -> clasimpset |
|
36 |
val addSDs2: clasimpset * thm list -> clasimpset |
|
37 |
val addIs2: clasimpset * thm list -> clasimpset |
|
38 |
val addEs2: clasimpset * thm list -> clasimpset |
|
39 |
val addDs2: clasimpset * thm list -> clasimpset |
|
40 |
val addsimps2: clasimpset * thm list -> clasimpset |
|
41 |
val delsimps2: clasimpset * thm list -> clasimpset |
|
42 |
val addSss: claset * simpset -> claset |
|
43 |
val addss: claset * simpset -> claset |
|
44 |
val addIffs: clasimpset * thm list -> clasimpset |
|
45 |
val delIffs: clasimpset * thm list -> clasimpset |
|
46 |
val AddIffs: thm list -> unit |
|
47 |
val DelIffs: thm list -> unit |
|
48 |
val CLASIMPSET: (clasimpset -> tactic) -> tactic |
|
49 |
val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic |
|
5483 | 50 |
val clarsimp_tac: clasimpset -> int -> tactic |
9860 | 51 |
val Clarsimp_tac: int -> tactic |
52 |
val mk_auto_tac: clasimpset -> int -> int -> tactic |
|
53 |
val auto_tac: clasimpset -> tactic |
|
54 |
val Auto_tac: tactic |
|
55 |
val auto: unit -> unit |
|
56 |
val force_tac: clasimpset -> int -> tactic |
|
57 |
val Force_tac: int -> tactic |
|
58 |
val force: int -> unit |
|
59 |
val fast_simp_tac: clasimpset -> int -> tactic |
|
60 |
val slow_simp_tac: clasimpset -> int -> tactic |
|
61 |
val best_simp_tac: clasimpset -> int -> tactic |
|
8639 | 62 |
val change_global_css: (clasimpset * thm list -> clasimpset) -> theory attribute |
63 |
val change_local_css: (clasimpset * thm list -> clasimpset) -> Proof.context attribute |
|
9506 | 64 |
val get_local_clasimpset: Proof.context -> clasimpset |
9860 | 65 |
val iff_add_global: theory attribute |
66 |
val iff_del_global: theory attribute |
|
67 |
val iff_add_local: Proof.context attribute |
|
68 |
val iff_del_local: Proof.context attribute |
|
69 |
val iff_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list |
|
9506 | 70 |
val clasimp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list |
9860 | 71 |
val setup: (theory -> theory) list |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
72 |
end; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
73 |
|
5219 | 74 |
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
|
75 |
struct |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
76 |
|
5219 | 77 |
open Data; |
78 |
||
79 |
type claset = Classical.claset; |
|
80 |
type simpset = Simplifier.simpset; |
|
81 |
type clasimpset = claset * simpset; |
|
82 |
||
83 |
||
84 |
(* clasimpset operations *) |
|
85 |
||
86 |
(*this interface for updating a clasimpset is rudimentary and just for |
|
87 |
convenience for the most common cases*) |
|
88 |
||
89 |
fun pair_upd1 f ((a,b),x) = (f(a,x), b); |
|
90 |
fun pair_upd2 f ((a,b),x) = (a, f(b,x)); |
|
91 |
||
92 |
fun op addSIs2 arg = pair_upd1 Classical.addSIs arg; |
|
93 |
fun op addSEs2 arg = pair_upd1 Classical.addSEs arg; |
|
94 |
fun op addSDs2 arg = pair_upd1 Classical.addSDs arg; |
|
95 |
fun op addIs2 arg = pair_upd1 Classical.addIs arg; |
|
96 |
fun op addEs2 arg = pair_upd1 Classical.addEs arg; |
|
97 |
fun op addDs2 arg = pair_upd1 Classical.addDs arg; |
|
98 |
fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg; |
|
99 |
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
|
100 |
|
9402 | 101 |
(*not totally safe: may instantiate unknowns that appear also in other subgoals*) |
102 |
val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true); |
|
103 |
||
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
104 |
(*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
|
105 |
(*Caution: only one simpset added can be added by each of addSss and addss*) |
5567 | 106 |
fun cs addSss ss = Classical.addSaltern (cs, ("safe_asm_full_simp_tac", |
9772 | 107 |
CHANGED o safe_asm_full_simp_tac ss)); |
108 |
fun cs addss ss = Classical.addbefore (cs, ("asm_full_simp_tac", |
|
109 |
CHANGED o Simplifier.asm_full_simp_tac ss)); |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
110 |
|
9860 | 111 |
|
112 |
(* iffs: addition of rules to simpsets and clasets simultaneously *) |
|
113 |
||
114 |
(*Takes UNCONDITIONAL theorems of the form A<->B to |
|
115 |
the Safe Intr rule B==>A and |
|
116 |
the Safe Destruct rule A==>B. |
|
117 |
Also ~A goes to the Safe Elim rule A ==> ?R |
|
118 |
Failing other cases, A is added as a Safe Intr rule*) |
|
119 |
local |
|
120 |
||
121 |
fun addIff ((cla, simp), th) = |
|
122 |
(case dest_Trueprop (#prop (rep_thm th)) of |
|
123 |
con $ _ $ _ => |
|
124 |
if con = Data.iff_const then |
|
125 |
Classical.addSDs (Classical.addSIs (cla, [zero_var_indexes (th RS Data.iffD2)]), |
|
126 |
[zero_var_indexes (th RS Data.iffD1)]) |
|
127 |
else Classical.addSIs (cla, [th]) |
|
128 |
| con $ A => |
|
129 |
if con = Data.not_const then Classical.addSEs (cla, [zero_var_indexes (th RS Data.notE)]) |
|
130 |
else Classical.addSIs (cla, [th]) |
|
131 |
| _ => Classical.addSIs (cla, [th]), Simplifier.addsimps (simp, [th])) |
|
132 |
handle TERM _ => error ("iff add: theorem must be unconditional\n" ^ Display.string_of_thm th); |
|
133 |
||
134 |
fun delIff ((cla, simp), th) = |
|
135 |
(case dest_Trueprop (#prop (rep_thm th)) of |
|
136 |
con $ _ $ _ => |
|
137 |
if con = Data.iff_const then |
|
138 |
Classical.delrules (cla, [zero_var_indexes (th RS Data.iffD2), |
|
139 |
Data.cla_make_elim (zero_var_indexes (th RS Data.iffD1))]) |
|
140 |
else Classical.delrules (cla, [th]) |
|
141 |
| con $ A => |
|
142 |
if con = Data.not_const then Classical.delrules (cla, [zero_var_indexes (th RS Data.notE)]) |
|
143 |
else Classical.delrules (cla, [th]) |
|
144 |
| _ => Classical.delrules (cla, [th]), Simplifier.delsimps (simp, [th])) |
|
145 |
handle TERM _ => |
|
146 |
(warning ("iff del: ignoring conditional theorem\n" ^ string_of_thm th); (cla, simp)); |
|
147 |
||
148 |
fun store_clasimp (cla, simp) = |
|
149 |
(Classical.claset_ref () := cla; Simplifier.simpset_ref () := simp); |
|
150 |
||
151 |
in |
|
152 |
||
153 |
val op addIffs = foldl addIff; |
|
154 |
val op delIffs = foldl delIff; |
|
155 |
fun AddIffs thms = store_clasimp ((Classical.claset (), Simplifier.simpset ()) addIffs thms); |
|
156 |
fun DelIffs thms = store_clasimp ((Classical.claset (), Simplifier.simpset ()) delIffs thms); |
|
157 |
||
158 |
end; |
|
159 |
||
160 |
||
5219 | 161 |
(* tacticals *) |
162 |
||
163 |
fun CLASIMPSET tacf state = |
|
164 |
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state; |
|
165 |
||
166 |
fun CLASIMPSET' tacf i state = |
|
167 |
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state; |
|
4888 | 168 |
|
169 |
||
9402 | 170 |
fun clarsimp_tac (cs, ss) = safe_asm_full_simp_tac ss THEN_ALL_NEW |
9772 | 171 |
Classical.clarify_tac (cs addSss ss); |
5483 | 172 |
fun Clarsimp_tac i = clarsimp_tac (Classical.claset(), Simplifier.simpset()) i; |
173 |
||
174 |
||
5219 | 175 |
(* auto_tac *) |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
176 |
|
5219 | 177 |
fun blast_depth_tac cs m i thm = |
9772 | 178 |
Blast.depth_tac cs m i thm |
5554 | 179 |
handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty); |
9772 | 180 |
|
181 |
(* a variant of depth_tac that avoids interference of the simplifier |
|
5219 | 182 |
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
|
183 |
local |
9772 | 184 |
fun slow_step_tac' cs = Classical.appWrappers cs |
185 |
(Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs); |
|
186 |
in fun nodup_depth_tac cs m i state = SELECT_GOAL |
|
187 |
(Classical.safe_steps_tac cs 1 THEN_ELSE |
|
188 |
(DEPTH_SOLVE (nodup_depth_tac cs m 1), |
|
189 |
Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac |
|
190 |
(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
|
191 |
)) i state; |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
192 |
end; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
193 |
|
9402 | 194 |
(*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
|
195 |
in some of the subgoals*) |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
196 |
fun mk_auto_tac (cs, ss) m n = |
5219 | 197 |
let val cs' = cs addss ss |
9772 | 198 |
val maintac = |
199 |
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
|
200 |
ORELSE' |
5756
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
201 |
(CHANGED o nodup_depth_tac cs' n); (* slower but more general *) |
5219 | 202 |
in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), |
9772 | 203 |
TRY (Classical.safe_tac cs), |
204 |
REPEAT (FIRSTGOAL maintac), |
|
5219 | 205 |
TRY (Classical.safe_tac (cs addSss ss)), |
9772 | 206 |
prune_params_tac] |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
207 |
end; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
208 |
|
9772 | 209 |
fun auto_tac css = mk_auto_tac css 4 2; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
210 |
|
5219 | 211 |
fun Auto_tac st = auto_tac (Classical.claset(), Simplifier.simpset()) st; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
212 |
|
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
213 |
fun auto () = by Auto_tac; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
214 |
|
9772 | 215 |
|
5219 | 216 |
(* force_tac *) |
217 |
||
4659 | 218 |
(* 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
|
219 |
fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [ |
9772 | 220 |
Classical.clarify_tac cs' 1, |
221 |
IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1), |
|
222 |
ALLGOALS (Classical.first_best_tac cs')]) end; |
|
5219 | 223 |
fun Force_tac i = force_tac (Classical.claset(), Simplifier.simpset()) i; |
4717
1370ad043564
renamed smart_tac to force_tac, slight improvement of force_tac
oheimb
parents:
4659
diff
changeset
|
224 |
fun force i = by (Force_tac i); |
4659 | 225 |
|
5219 | 226 |
|
9805 | 227 |
(* basic combinations *) |
228 |
||
229 |
fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end; |
|
9591 | 230 |
|
9805 | 231 |
val fast_simp_tac = ADDSS Classical.fast_tac; |
232 |
val slow_simp_tac = ADDSS Classical.slow_tac; |
|
233 |
val best_simp_tac = ADDSS Classical.best_tac; |
|
9591 | 234 |
|
235 |
||
9860 | 236 |
(* access clasimpset *) |
8639 | 237 |
|
238 |
fun change_global_css f (thy, th) = |
|
239 |
let |
|
240 |
val cs_ref = Classical.claset_ref_of thy; |
|
241 |
val ss_ref = Simplifier.simpset_ref_of thy; |
|
242 |
val (cs', ss') = f ((! cs_ref, ! ss_ref), [th]); |
|
243 |
in cs_ref := cs'; ss_ref := ss'; (thy, th) end; |
|
244 |
||
245 |
fun change_local_css f (ctxt, th) = |
|
246 |
let |
|
247 |
val cs = Classical.get_local_claset ctxt; |
|
248 |
val ss = Simplifier.get_local_simpset ctxt; |
|
249 |
val (cs', ss') = f ((cs, ss), [th]); |
|
250 |
val ctxt' = |
|
251 |
ctxt |
|
252 |
|> Classical.put_local_claset cs' |
|
253 |
|> Simplifier.put_local_simpset ss'; |
|
254 |
in (ctxt', th) end; |
|
255 |
||
7153 | 256 |
fun get_local_clasimpset ctxt = |
257 |
(Classical.get_local_claset ctxt, Simplifier.get_local_simpset ctxt); |
|
5926 | 258 |
|
9860 | 259 |
|
260 |
(* attributes *) |
|
261 |
||
262 |
val iff_add_global = change_global_css (op addIffs); |
|
263 |
val iff_del_global = change_global_css (op delIffs); |
|
264 |
val iff_add_local = change_local_css (op addIffs); |
|
265 |
val iff_del_local = change_local_css (op delIffs); |
|
266 |
||
267 |
val iff_attr = |
|
268 |
(Attrib.add_del_args iff_add_global iff_del_global, |
|
269 |
Attrib.add_del_args iff_add_local iff_del_local); |
|
270 |
||
271 |
||
272 |
(* method modifiers *) |
|
273 |
||
274 |
val iffN = "iff"; |
|
275 |
val addN = "add"; |
|
276 |
val delN = "del"; |
|
277 |
||
278 |
val iff_modifiers = |
|
279 |
[Args.$$$ iffN -- Args.colon >> K ((I, iff_add_local): Method.modifier), |
|
280 |
Args.$$$ iffN -- Args.$$$ addN -- Args.colon >> K (I, iff_add_local), |
|
281 |
Args.$$$ iffN -- Args.$$$ delN -- Args.colon >> K (I, iff_del_local)]; |
|
282 |
||
8469 | 283 |
val clasimp_modifiers = |
9860 | 284 |
Simplifier.simp_modifiers @ Splitter.split_modifiers @ |
285 |
Classical.cla_modifiers @ iff_modifiers; |
|
286 |
||
287 |
||
288 |
(* methods *) |
|
5926 | 289 |
|
7559 | 290 |
fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts => |
291 |
ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (get_local_clasimpset ctxt)); |
|
7132 | 292 |
|
7559 | 293 |
fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts => |
8168 | 294 |
HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (get_local_clasimpset ctxt))); |
5926 | 295 |
|
7559 | 296 |
val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth; |
297 |
val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth'; |
|
5926 | 298 |
|
299 |
||
9772 | 300 |
fun auto_args m = |
301 |
Method.bang_sectioned_args' clasimp_modifiers (Scan.lift (Scan.option (Args.nat -- Args.nat))) m; |
|
302 |
||
303 |
fun auto_meth None = clasimp_meth (CHANGED o auto_tac) |
|
304 |
| auto_meth (Some (m, n)) = clasimp_meth (CHANGED o (fn css => mk_auto_tac css m n)); |
|
305 |
||
306 |
||
307 |
(* theory setup *) |
|
308 |
||
5926 | 309 |
val setup = |
9860 | 310 |
[Attrib.add_attributes |
9893 | 311 |
[("iff", iff_attr, "declaration of Simplifier / Classical rules")], |
9860 | 312 |
Method.add_methods |
9591 | 313 |
[("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"), |
9805 | 314 |
("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"), |
315 |
("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"), |
|
9591 | 316 |
("force", clasimp_method' force_tac, "force"), |
9772 | 317 |
("auto", auto_args auto_meth, "auto"), |
9591 | 318 |
("clarsimp", clasimp_method' (CHANGED oo clarsimp_tac), "clarify simplified goal")]]; |
5926 | 319 |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
320 |
end; |