author | paulson |
Thu, 25 May 2000 15:10:57 +0200 | |
changeset 8967 | 00f18476ac15 |
parent 8639 | 31bcb6b64d60 |
child 9402 | 480a1b40fdd6 |
permissions | -rw-r--r-- |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
1 |
(* Title: Provers/clasimp.ML |
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 |
|
5554 | 9 |
infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 addsimps2 delsimps2; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
10 |
infix 4 addSss addss; |
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 |
|
19 |
end |
|
20 |
||
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
21 |
signature CLASIMP = |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
22 |
sig |
5219 | 23 |
include CLASIMP_DATA |
24 |
type claset |
|
25 |
type simpset |
|
26 |
type clasimpset |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
27 |
val addSIs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
28 |
val addSEs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
29 |
val addSDs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
30 |
val addIs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
31 |
val addEs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
32 |
val addDs2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
33 |
val addsimps2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
34 |
val delsimps2 : clasimpset * thm list -> clasimpset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
35 |
val addSss : claset * simpset -> claset |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
36 |
val addss : claset * simpset -> claset |
5483 | 37 |
val CLASIMPSET :(clasimpset -> tactic) -> tactic |
38 |
val CLASIMPSET' :(clasimpset -> 'a -> tactic) -> 'a -> tactic |
|
39 |
val clarsimp_tac: clasimpset -> int -> tactic |
|
40 |
val Clarsimp_tac: int -> tactic |
|
41 |
val mk_auto_tac : clasimpset -> int -> int -> tactic |
|
42 |
val auto_tac : clasimpset -> tactic |
|
43 |
val Auto_tac : tactic |
|
44 |
val auto : unit -> unit |
|
45 |
val force_tac : clasimpset -> int -> tactic |
|
46 |
val Force_tac : int -> tactic |
|
47 |
val force : int -> unit |
|
8639 | 48 |
val change_global_css: (clasimpset * thm list -> clasimpset) -> theory attribute |
49 |
val change_local_css: (clasimpset * thm list -> clasimpset) -> Proof.context attribute |
|
5926 | 50 |
val setup : (theory -> theory) list |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
51 |
end; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
52 |
|
5219 | 53 |
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
|
54 |
struct |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
55 |
|
5219 | 56 |
open Data; |
57 |
||
58 |
type claset = Classical.claset; |
|
59 |
type simpset = Simplifier.simpset; |
|
60 |
type clasimpset = claset * simpset; |
|
61 |
||
62 |
||
63 |
(* clasimpset operations *) |
|
64 |
||
65 |
(*this interface for updating a clasimpset is rudimentary and just for |
|
66 |
convenience for the most common cases*) |
|
67 |
||
68 |
fun pair_upd1 f ((a,b),x) = (f(a,x), b); |
|
69 |
fun pair_upd2 f ((a,b),x) = (a, f(b,x)); |
|
70 |
||
71 |
fun op addSIs2 arg = pair_upd1 Classical.addSIs arg; |
|
72 |
fun op addSEs2 arg = pair_upd1 Classical.addSEs arg; |
|
73 |
fun op addSDs2 arg = pair_upd1 Classical.addSDs arg; |
|
74 |
fun op addIs2 arg = pair_upd1 Classical.addIs arg; |
|
75 |
fun op addEs2 arg = pair_upd1 Classical.addEs arg; |
|
76 |
fun op addDs2 arg = pair_upd1 Classical.addDs arg; |
|
77 |
fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg; |
|
78 |
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
|
79 |
|
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
80 |
(*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
|
81 |
(*Caution: only one simpset added can be added by each of addSss and addss*) |
5567 | 82 |
fun cs addSss ss = Classical.addSaltern (cs, ("safe_asm_full_simp_tac", |
5554 | 83 |
CHANGED o Simplifier.safe_asm_full_simp_tac ss)); |
5567 | 84 |
fun cs addss ss = Classical.addbefore (cs, ("asm_full_simp_tac", |
5554 | 85 |
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
|
86 |
|
5219 | 87 |
(* tacticals *) |
88 |
||
89 |
fun CLASIMPSET tacf state = |
|
90 |
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state; |
|
91 |
||
92 |
fun CLASIMPSET' tacf i state = |
|
93 |
Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state; |
|
4888 | 94 |
|
95 |
||
7957
0196b2302e21
clarsimp_tac now copes with the (unwanted) case that the simplifier splits
oheimb
parents:
7559
diff
changeset
|
96 |
fun clarsimp_tac (cs, ss) = Simplifier.safe_asm_full_simp_tac ss THEN_ALL_NEW |
5483 | 97 |
Classical.clarify_tac (cs addSss ss); |
98 |
fun Clarsimp_tac i = clarsimp_tac (Classical.claset(), Simplifier.simpset()) i; |
|
99 |
||
100 |
||
5219 | 101 |
(* auto_tac *) |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
102 |
|
5219 | 103 |
fun blast_depth_tac cs m i thm = |
5554 | 104 |
Blast.depth_tac cs m i thm |
105 |
handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty); |
|
5756
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
106 |
|
5219 | 107 |
(* a variant of depth_tac that avoids interference of the simplifier |
108 |
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
|
109 |
local |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
110 |
fun slow_step_tac' cs = Classical.appWrappers cs |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
111 |
(Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs); |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
112 |
in fun nodup_depth_tac cs m i state = SELECT_GOAL |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
113 |
(Classical.safe_steps_tac cs 1 THEN_ELSE |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
114 |
(DEPTH_SOLVE (nodup_depth_tac cs m 1), |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
115 |
Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
116 |
(slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m-1) 1)) |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
117 |
)) i state; |
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
118 |
end; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
119 |
|
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
120 |
(*Designed to be idempotent, except if best_tac instantiates variables |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
121 |
in some of the subgoals*) |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
122 |
fun mk_auto_tac (cs, ss) m n = |
5219 | 123 |
let val cs' = cs addss ss |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
124 |
val maintac = |
5756
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
125 |
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
|
126 |
ORELSE' |
5756
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
127 |
(CHANGED o nodup_depth_tac cs' n); (* slower but more general *) |
5219 | 128 |
in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), |
5525 | 129 |
TRY (Classical.safe_tac cs), |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
130 |
REPEAT (FIRSTGOAL maintac), |
5219 | 131 |
TRY (Classical.safe_tac (cs addSss ss)), |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
132 |
prune_params_tac] |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
133 |
end; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
134 |
|
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
135 |
fun auto_tac (cs,ss) = mk_auto_tac (cs,ss) 4 2; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
136 |
|
5219 | 137 |
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
|
138 |
|
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
139 |
fun auto () = by Auto_tac; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
140 |
|
5219 | 141 |
|
142 |
(* force_tac *) |
|
143 |
||
4659 | 144 |
(* 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
|
145 |
fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [ |
5219 | 146 |
Classical.clarify_tac cs' 1, |
147 |
IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1), |
|
148 |
ALLGOALS (Classical.best_tac cs')]) end; |
|
149 |
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
|
150 |
fun force i = by (Force_tac i); |
4659 | 151 |
|
5219 | 152 |
|
8639 | 153 |
(* attributes *) |
154 |
||
155 |
fun change_global_css f (thy, th) = |
|
156 |
let |
|
157 |
val cs_ref = Classical.claset_ref_of thy; |
|
158 |
val ss_ref = Simplifier.simpset_ref_of thy; |
|
159 |
val (cs', ss') = f ((! cs_ref, ! ss_ref), [th]); |
|
160 |
in cs_ref := cs'; ss_ref := ss'; (thy, th) end; |
|
161 |
||
162 |
fun change_local_css f (ctxt, th) = |
|
163 |
let |
|
164 |
val cs = Classical.get_local_claset ctxt; |
|
165 |
val ss = Simplifier.get_local_simpset ctxt; |
|
166 |
val (cs', ss') = f ((cs, ss), [th]); |
|
167 |
val ctxt' = |
|
168 |
ctxt |
|
169 |
|> Classical.put_local_claset cs' |
|
170 |
|> Simplifier.put_local_simpset ss'; |
|
171 |
in (ctxt', th) end; |
|
172 |
||
173 |
||
5926 | 174 |
(* methods *) |
175 |
||
7153 | 176 |
fun get_local_clasimpset ctxt = |
177 |
(Classical.get_local_claset ctxt, Simplifier.get_local_simpset ctxt); |
|
5926 | 178 |
|
8469 | 179 |
val clasimp_modifiers = |
180 |
Simplifier.simp_modifiers @ Splitter.split_modifiers @ Classical.cla_modifiers; |
|
5926 | 181 |
|
7559 | 182 |
fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts => |
183 |
ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (get_local_clasimpset ctxt)); |
|
7132 | 184 |
|
7559 | 185 |
fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts => |
8168 | 186 |
HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (get_local_clasimpset ctxt))); |
5926 | 187 |
|
7559 | 188 |
val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth; |
189 |
val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth'; |
|
5926 | 190 |
|
191 |
||
192 |
val setup = |
|
193 |
[Method.add_methods |
|
7437
0f99a2103ea0
renamed improper method 'clarsimp' to 'clarsimp_tac';
wenzelm
parents:
7424
diff
changeset
|
194 |
[("clarsimp_tac", clasimp_method' clarsimp_tac, "clarsimp (improper!)"), |
7394 | 195 |
("auto", clasimp_method (CHANGED o auto_tac), "auto"), |
7272 | 196 |
("force", clasimp_method' force_tac, "force")]]; |
197 |
||
5926 | 198 |
|
199 |
||
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
200 |
end; |