author | wenzelm |
Sat, 30 May 2015 23:30:54 +0200 | |
changeset 60318 | ab785001b51d |
parent 59582 | 0fbed69ff081 |
child 60366 | df3e916bcd26 |
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 |
Author: David von Oheimb, TU Muenchen |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
3 |
|
5219 | 4 |
Combination of classical reasoner and simplifier (depends on |
16019 | 5 |
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
|
6 |
*) |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
7 |
|
5219 | 8 |
signature CLASIMP_DATA = |
9 |
sig |
|
8469 | 10 |
structure Splitter: SPLITTER |
5219 | 11 |
structure Classical: CLASSICAL |
12 |
structure Blast: BLAST |
|
9860 | 13 |
val notE: thm |
14 |
val iffD1: thm |
|
15 |
val iffD2: thm |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
16 |
end; |
5219 | 17 |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
18 |
signature CLASIMP = |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
19 |
sig |
42793 | 20 |
val addSss: Proof.context -> Proof.context |
21 |
val addss: Proof.context -> Proof.context |
|
22 |
val clarsimp_tac: Proof.context -> int -> tactic |
|
23 |
val mk_auto_tac: Proof.context -> int -> int -> tactic |
|
24 |
val auto_tac: Proof.context -> tactic |
|
25 |
val force_tac: Proof.context -> int -> tactic |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43331
diff
changeset
|
26 |
val fast_force_tac: Proof.context -> int -> tactic |
42793 | 27 |
val slow_simp_tac: Proof.context -> int -> tactic |
28 |
val best_simp_tac: Proof.context -> int -> tactic |
|
18728 | 29 |
val iff_add: attribute |
30 |
val iff_add': attribute |
|
31 |
val iff_del: attribute |
|
30513 | 32 |
val iff_modifiers: Method.modifier parser list |
33 |
val clasimp_modifiers: Method.modifier parser list |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
34 |
end; |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
35 |
|
42478 | 36 |
functor Clasimp(Data: CLASIMP_DATA): CLASIMP = |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
37 |
struct |
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
38 |
|
42478 | 39 |
structure Splitter = Data.Splitter; |
40 |
structure Classical = Data.Classical; |
|
41 |
structure Blast = Data.Blast; |
|
5219 | 42 |
|
26497
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
wenzelm
parents:
26425
diff
changeset
|
43 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
44 |
(* simp as classical wrapper *) |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
45 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
46 |
(* FIXME !? *) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
47 |
fun clasimp f name tac ctxt = f (ctxt, (name, fn _ => CHANGED o tac ctxt)); |
42478 | 48 |
|
42793 | 49 |
(*Add a simpset to the claset!*) |
50 |
(*Caution: only one simpset added can be added by each of addSss and addss*) |
|
50111 | 51 |
val addSss = clasimp Classical.addSafter "safe_asm_full_simp_tac" Simplifier.safe_asm_full_simp_tac; |
42793 | 52 |
val addss = clasimp Classical.addbefore "asm_full_simp_tac" Simplifier.asm_full_simp_tac; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
53 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
54 |
|
9860 | 55 |
(* iffs: addition of rules to simpsets and clasets simultaneously *) |
56 |
||
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
57 |
local |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
58 |
|
11344 | 59 |
(*Takes (possibly conditional) theorems of the form A<->B to |
9860 | 60 |
the Safe Intr rule B==>A and |
61 |
the Safe Destruct rule A==>B. |
|
62 |
Also ~A goes to the Safe Elim rule A ==> ?R |
|
11462 | 63 |
Failing other cases, A is added as a Safe Intr rule*) |
9860 | 64 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
65 |
fun add_iff safe unsafe = |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
66 |
Thm.declaration_attribute (fn th => |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
67 |
let |
59582 | 68 |
val n = Thm.nprems_of th; |
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
69 |
val (elim, intro) = if n = 0 then safe else unsafe; |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
70 |
val zero_rotate = zero_var_indexes o rotate_prems n; |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
71 |
in |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
72 |
Thm.attribute_declaration intro (zero_rotate (th RS Data.iffD2)) #> |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
73 |
Thm.attribute_declaration elim (Tactic.make_elim (zero_rotate (th RS Data.iffD1))) |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
74 |
handle THM _ => |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
75 |
(Thm.attribute_declaration elim (zero_rotate (th RS Data.notE)) |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
76 |
handle THM _ => Thm.attribute_declaration intro th) |
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
77 |
end); |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
78 |
|
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
79 |
fun del_iff del = Thm.declaration_attribute (fn th => |
59582 | 80 |
let val zero_rotate = zero_var_indexes o rotate_prems (Thm.nprems_of th) in |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
81 |
Thm.attribute_declaration del (zero_rotate (th RS Data.iffD2)) #> |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
82 |
Thm.attribute_declaration del (Tactic.make_elim (zero_rotate (th RS Data.iffD1))) |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
83 |
handle THM _ => |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
84 |
(Thm.attribute_declaration del (zero_rotate (th RS Data.notE)) |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
85 |
handle THM _ => Thm.attribute_declaration del th) |
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
86 |
end); |
18630 | 87 |
|
9860 | 88 |
in |
89 |
||
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
90 |
val iff_add = |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
91 |
Thm.declaration_attribute (fn th => |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
92 |
Thm.attribute_declaration (add_iff |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
93 |
(Classical.safe_elim NONE, Classical.safe_intro NONE) |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
94 |
(Classical.haz_elim NONE, Classical.haz_intro NONE)) th |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
95 |
#> Thm.attribute_declaration Simplifier.simp_add th); |
10033
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
wenzelm
parents:
9952
diff
changeset
|
96 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
97 |
val iff_add' = |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
98 |
add_iff |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
99 |
(Context_Rules.elim_query NONE, Context_Rules.intro_query NONE) |
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
100 |
(Context_Rules.elim_query NONE, Context_Rules.intro_query NONE); |
12375 | 101 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
102 |
val iff_del = |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
103 |
Thm.declaration_attribute (fn th => |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
104 |
Thm.attribute_declaration (del_iff Classical.rule_del) th #> |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
105 |
Thm.attribute_declaration (del_iff Context_Rules.rule_del) th #> |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
44890
diff
changeset
|
106 |
Thm.attribute_declaration Simplifier.simp_del th); |
12375 | 107 |
|
9860 | 108 |
end; |
109 |
||
110 |
||
42478 | 111 |
(* tactics *) |
5219 | 112 |
|
42793 | 113 |
fun clarsimp_tac ctxt = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
114 |
Simplifier.safe_asm_full_simp_tac ctxt THEN_ALL_NEW |
42793 | 115 |
Classical.clarify_tac (addSss ctxt); |
12375 | 116 |
|
5483 | 117 |
|
5219 | 118 |
(* auto_tac *) |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
119 |
|
9772 | 120 |
(* a variant of depth_tac that avoids interference of the simplifier |
5219 | 121 |
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
|
122 |
local |
42478 | 123 |
|
42793 | 124 |
fun slow_step_tac' ctxt = |
125 |
Classical.appWrappers ctxt |
|
126 |
(Classical.instp_step_tac ctxt APPEND' Classical.haz_step_tac ctxt); |
|
42478 | 127 |
|
128 |
in |
|
129 |
||
42793 | 130 |
fun nodup_depth_tac ctxt m i st = |
42478 | 131 |
SELECT_GOAL |
42793 | 132 |
(Classical.safe_steps_tac ctxt 1 THEN_ELSE |
133 |
(DEPTH_SOLVE (nodup_depth_tac ctxt m 1), |
|
134 |
Classical.inst0_step_tac ctxt 1 APPEND COND (K (m = 0)) no_tac |
|
135 |
(slow_step_tac' ctxt 1 THEN DEPTH_SOLVE (nodup_depth_tac ctxt (m - 1) 1)))) i st; |
|
42479 | 136 |
|
5756
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5567
diff
changeset
|
137 |
end; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
138 |
|
43331
01f051619eee
clarified raw_blast, which is not really a tactic since it operates directly on subgoal 1 without bounds check (cf. c46107e6714b);
wenzelm
parents:
42805
diff
changeset
|
139 |
(*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
|
140 |
in some of the subgoals*) |
42793 | 141 |
fun mk_auto_tac ctxt m n = |
42478 | 142 |
let |
143 |
val main_tac = |
|
43331
01f051619eee
clarified raw_blast, which is not really a tactic since it operates directly on subgoal 1 without bounds check (cf. c46107e6714b);
wenzelm
parents:
42805
diff
changeset
|
144 |
Blast.depth_tac ctxt m (* fast but can't use wrappers *) |
42478 | 145 |
ORELSE' |
42793 | 146 |
(CHANGED o nodup_depth_tac (addss ctxt) n); (* slower but more general *) |
42478 | 147 |
in |
58008 | 148 |
PARALLEL_ALLGOALS (Simplifier.asm_full_simp_tac ctxt) THEN |
42793 | 149 |
TRY (Classical.safe_tac ctxt) THEN |
42479 | 150 |
REPEAT_DETERM (FIRSTGOAL main_tac) THEN |
42793 | 151 |
TRY (Classical.safe_tac (addSss ctxt)) THEN |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
152 |
prune_params_tac ctxt |
42478 | 153 |
end; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
154 |
|
42793 | 155 |
fun auto_tac ctxt = mk_auto_tac ctxt 4 2; |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
156 |
|
9772 | 157 |
|
5219 | 158 |
(* force_tac *) |
159 |
||
4659 | 160 |
(* aimed to solve the given subgoal totally, using whatever tools possible *) |
42793 | 161 |
fun force_tac ctxt = |
162 |
let val ctxt' = addss ctxt in |
|
42479 | 163 |
SELECT_GOAL |
42793 | 164 |
(Classical.clarify_tac ctxt' 1 THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
165 |
IF_UNSOLVED (Simplifier.asm_full_simp_tac ctxt 1) THEN |
42793 | 166 |
ALLGOALS (Classical.first_best_tac ctxt')) |
42478 | 167 |
end; |
4659 | 168 |
|
5219 | 169 |
|
9805 | 170 |
(* basic combinations *) |
171 |
||
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43331
diff
changeset
|
172 |
val fast_force_tac = Classical.fast_tac o addss; |
42793 | 173 |
val slow_simp_tac = Classical.slow_tac o addss; |
174 |
val best_simp_tac = Classical.best_tac o addss; |
|
9591 | 175 |
|
47967
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
45375
diff
changeset
|
176 |
|
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
45375
diff
changeset
|
177 |
|
42784
a2dca9a3d0da
simplified clasimpset declarations -- prefer attributes;
wenzelm
parents:
42479
diff
changeset
|
178 |
(** concrete syntax **) |
9860 | 179 |
|
180 |
(* attributes *) |
|
181 |
||
58826 | 182 |
val _ = |
183 |
Theory.setup |
|
184 |
(Attrib.setup @{binding iff} |
|
185 |
(Scan.lift |
|
186 |
(Args.del >> K iff_del || |
|
187 |
Scan.option Args.add -- Args.query >> K iff_add' || |
|
188 |
Scan.option Args.add >> K iff_add)) |
|
189 |
"declaration of Simplifier / Classical rules"); |
|
9860 | 190 |
|
191 |
||
192 |
(* method modifiers *) |
|
193 |
||
194 |
val iffN = "iff"; |
|
195 |
||
196 |
val iff_modifiers = |
|
58048
aa6296d09e0e
more explicit Method.modifier with reported position;
wenzelm
parents:
58008
diff
changeset
|
197 |
[Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K (Method.modifier iff_add @{here}), |
aa6296d09e0e
more explicit Method.modifier with reported position;
wenzelm
parents:
58008
diff
changeset
|
198 |
Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (Method.modifier iff_add' @{here}), |
aa6296d09e0e
more explicit Method.modifier with reported position;
wenzelm
parents:
58008
diff
changeset
|
199 |
Args.$$$ iffN -- Args.del -- Args.colon >> K (Method.modifier iff_del @{here})]; |
9860 | 200 |
|
8469 | 201 |
val clasimp_modifiers = |
9860 | 202 |
Simplifier.simp_modifiers @ Splitter.split_modifiers @ |
203 |
Classical.cla_modifiers @ iff_modifiers; |
|
204 |
||
205 |
||
206 |
(* methods *) |
|
5926 | 207 |
|
30541 | 208 |
fun clasimp_method' tac = |
42793 | 209 |
Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o tac); |
9772 | 210 |
|
30541 | 211 |
val auto_method = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36601
diff
changeset
|
212 |
Scan.lift (Scan.option (Parse.nat -- Parse.nat)) --| |
35613 | 213 |
Method.sections clasimp_modifiers >> |
42793 | 214 |
(fn NONE => SIMPLE_METHOD o CHANGED_PROP o auto_tac |
215 |
| SOME (m, n) => (fn ctxt => SIMPLE_METHOD (CHANGED_PROP (mk_auto_tac ctxt m n)))); |
|
9772 | 216 |
|
58826 | 217 |
val _ = |
218 |
Theory.setup |
|
219 |
(Method.setup @{binding fastforce} (clasimp_method' fast_force_tac) "combined fast and simp" #> |
|
220 |
Method.setup @{binding slowsimp} (clasimp_method' slow_simp_tac) "combined slow and simp" #> |
|
221 |
Method.setup @{binding bestsimp} (clasimp_method' best_simp_tac) "combined best and simp" #> |
|
222 |
Method.setup @{binding force} (clasimp_method' force_tac) "force" #> |
|
223 |
Method.setup @{binding auto} auto_method "auto" #> |
|
224 |
Method.setup @{binding clarsimp} (clasimp_method' (CHANGED_PROP oo clarsimp_tac)) |
|
225 |
"clarify simplified goal"); |
|
5926 | 226 |
|
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff
changeset
|
227 |
end; |