author | kleing |
Mon, 29 Nov 2004 06:09:45 +0100 | |
changeset 15336 | cb35ae957c65 |
parent 15036 | cab1c1fc1851 |
child 15452 | e2a721567f67 |
permissions | -rw-r--r-- |
1243 | 1 |
(* Title: Provers/simplifier.ML |
1 | 2 |
ID: $Id$ |
3557 | 3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
1 | 4 |
|
10412
1a1b4c1b2b7c
Moved meta simplification stuff from Thm to MetaSimplifier.
berghofe
parents:
10034
diff
changeset
|
5 |
Generic simplifier, suitable for most logics. See Pure/meta_simplifier.ML |
1a1b4c1b2b7c
Moved meta simplification stuff from Thm to MetaSimplifier.
berghofe
parents:
10034
diff
changeset
|
6 |
for the actual meta-level rewriting engine. |
1 | 7 |
*) |
1260 | 8 |
|
4795 | 9 |
signature BASIC_SIMPLIFIER = |
0 | 10 |
sig |
15024 | 11 |
include BASIC_META_SIMPLIFIER |
15036 | 12 |
type context_solver |
13 |
val mk_context_solver: string -> (Proof.context -> thm list -> int -> tactic) |
|
14 |
-> context_solver |
|
15 |
type context_simproc |
|
16 |
val mk_context_simproc: string -> cterm list -> |
|
17 |
(Proof.context -> simpset -> term -> thm option) -> context_simproc |
|
4366 | 18 |
val print_simpset: theory -> unit |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
19 |
val simpset_ref_of_sg: Sign.sg -> simpset ref |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
20 |
val simpset_ref_of: theory -> simpset ref |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
21 |
val simpset_of_sg: Sign.sg -> simpset |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
22 |
val simpset_of: theory -> simpset |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
23 |
val SIMPSET: (simpset -> tactic) -> tactic |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
24 |
val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
25 |
val simpset: unit -> simpset |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
26 |
val simpset_ref: unit -> simpset ref |
1243 | 27 |
val Addsimps: thm list -> unit |
28 |
val Delsimps: thm list -> unit |
|
2509 | 29 |
val Addsimprocs: simproc list -> unit |
30 |
val Delsimprocs: simproc list -> unit |
|
9715 | 31 |
val Addcongs: thm list -> unit |
32 |
val Delcongs: thm list -> unit |
|
15036 | 33 |
val local_simpset_of: Proof.context -> simpset |
10004 | 34 |
val safe_asm_full_simp_tac: simpset -> int -> tactic |
2629 | 35 |
val simp_tac: simpset -> int -> tactic |
36 |
val asm_simp_tac: simpset -> int -> tactic |
|
37 |
val full_simp_tac: simpset -> int -> tactic |
|
4722
d2e44673c21e
The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents:
4713
diff
changeset
|
38 |
val asm_lr_simp_tac: simpset -> int -> tactic |
2629 | 39 |
val asm_full_simp_tac: simpset -> int -> tactic |
40 |
val Simp_tac: int -> tactic |
|
41 |
val Asm_simp_tac: int -> tactic |
|
42 |
val Full_simp_tac: int -> tactic |
|
4722
d2e44673c21e
The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents:
4713
diff
changeset
|
43 |
val Asm_lr_simp_tac: int -> tactic |
2629 | 44 |
val Asm_full_simp_tac: int -> tactic |
3557 | 45 |
val simplify: simpset -> thm -> thm |
46 |
val asm_simplify: simpset -> thm -> thm |
|
47 |
val full_simplify: simpset -> thm -> thm |
|
14814
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
48 |
val asm_lr_simplify: simpset -> thm -> thm |
3557 | 49 |
val asm_full_simplify: simpset -> thm -> thm |
0 | 50 |
end; |
51 |
||
4795 | 52 |
signature SIMPLIFIER = |
53 |
sig |
|
54 |
include BASIC_SIMPLIFIER |
|
15036 | 55 |
val simproc_i: Sign.sg -> string -> term list |
56 |
-> (Sign.sg -> simpset -> term -> thm option) -> simproc |
|
57 |
val simproc: Sign.sg -> string -> string list |
|
58 |
-> (Sign.sg -> simpset -> term -> thm option) -> simproc |
|
59 |
val context_simproc_i: Sign.sg -> string -> term list |
|
60 |
-> (Proof.context -> simpset -> term -> thm option) -> context_simproc |
|
61 |
val context_simproc: Sign.sg -> string -> string list |
|
62 |
-> (Proof.context -> simpset -> term -> thm option) -> context_simproc |
|
5082 | 63 |
val rewrite: simpset -> cterm -> thm |
64 |
val asm_rewrite: simpset -> cterm -> thm |
|
65 |
val full_rewrite: simpset -> cterm -> thm |
|
14814
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
66 |
val asm_lr_rewrite: simpset -> cterm -> thm |
5082 | 67 |
val asm_full_rewrite: simpset -> cterm -> thm |
15036 | 68 |
val add_context_simprocs: context_simproc list -> theory -> theory |
69 |
val del_context_simprocs: context_simproc list -> theory -> theory |
|
70 |
val set_context_subgoaler: (Proof.context -> simpset -> int -> tactic) -> theory -> theory |
|
71 |
val reset_context_subgoaler: theory -> theory |
|
72 |
val add_context_looper: string * (Proof.context -> int -> Tactical.tactic) -> |
|
73 |
theory -> theory |
|
74 |
val del_context_looper: string -> theory -> theory |
|
75 |
val add_context_unsafe_solver: context_solver -> theory -> theory |
|
76 |
val add_context_safe_solver: context_solver -> theory -> theory |
|
5842 | 77 |
val print_local_simpset: Proof.context -> unit |
78 |
val get_local_simpset: Proof.context -> simpset |
|
79 |
val put_local_simpset: simpset -> Proof.context -> Proof.context |
|
8467 | 80 |
val change_global_ss: (simpset * thm list -> simpset) -> theory attribute |
81 |
val change_local_ss: (simpset * thm list -> simpset) -> Proof.context attribute |
|
4855 | 82 |
val simp_add_global: theory attribute |
83 |
val simp_del_global: theory attribute |
|
5842 | 84 |
val simp_add_local: Proof.context attribute |
85 |
val simp_del_local: Proof.context attribute |
|
9715 | 86 |
val cong_add_global: theory attribute |
87 |
val cong_del_global: theory attribute |
|
88 |
val cong_add_local: Proof.context attribute |
|
89 |
val cong_del_local: Proof.context attribute |
|
7177 | 90 |
val change_simpset_of: (simpset * 'a -> simpset) -> 'a -> theory -> theory |
7273 | 91 |
val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list |
5842 | 92 |
val setup: (theory -> theory) list |
8467 | 93 |
val method_setup: (Args.T list -> (Method.modifier * Args.T list)) list |
94 |
-> (theory -> theory) list |
|
8228 | 95 |
val easy_setup: thm -> thm list -> (theory -> theory) list |
4795 | 96 |
end; |
2503 | 97 |
|
98 |
structure Simplifier: SIMPLIFIER = |
|
0 | 99 |
struct |
100 |
||
15006
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents:
14814
diff
changeset
|
101 |
open MetaSimplifier; |
8467 | 102 |
|
15024 | 103 |
|
15036 | 104 |
(** context dependent simpset components **) |
105 |
||
106 |
(* datatype context_solver *) |
|
107 |
||
108 |
datatype context_solver = |
|
109 |
ContextSolver of (string * (Proof.context -> thm list -> int -> tactic)) * stamp; |
|
110 |
||
111 |
fun mk_context_solver name f = ContextSolver ((name, f), stamp ()); |
|
112 |
fun eq_context_solver (ContextSolver (_, id1), ContextSolver (_, id2)) = (id1 = id2); |
|
113 |
val merge_context_solvers = gen_merge_lists eq_context_solver; |
|
114 |
||
115 |
||
116 |
(* datatype context_simproc *) |
|
117 |
||
118 |
datatype context_simproc = ContextSimproc of |
|
119 |
(string * cterm list * (Proof.context -> simpset -> term -> thm option)) * stamp; |
|
120 |
||
121 |
fun mk_context_simproc name lhss f = ContextSimproc ((name, lhss, f), stamp ()); |
|
122 |
fun eq_context_simproc (ContextSimproc (_, id1), ContextSimproc (_, id2)) = (id1 = id2); |
|
123 |
val merge_context_simprocs = gen_merge_lists eq_context_simproc; |
|
124 |
||
125 |
fun context_simproc_i sg name = |
|
126 |
mk_context_simproc name o map (Thm.cterm_of sg o Logic.varify); |
|
127 |
||
128 |
fun context_simproc sg name = |
|
129 |
context_simproc_i sg name o map (Sign.simple_read_term sg TypeInfer.logicT); |
|
130 |
||
131 |
||
132 |
(* datatype context_ss *) |
|
133 |
||
134 |
datatype context_ss = ContextSS of |
|
135 |
{simprocs: context_simproc list, |
|
136 |
subgoal_tac: (Proof.context -> simpset -> int -> tactic) option, |
|
137 |
loop_tacs: (string * (Proof.context -> int -> tactic)) list, |
|
138 |
unsafe_solvers: context_solver list, |
|
139 |
solvers: context_solver list}; |
|
140 |
||
141 |
fun context_ss ctxt ss ctxt_ss = |
|
142 |
let |
|
143 |
val ContextSS {simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers} = ctxt_ss; |
|
144 |
fun prep_simproc (ContextSimproc ((name, lhss, f), _)) = |
|
145 |
mk_simproc name lhss (K (f ctxt)); |
|
146 |
fun add_loop (name, f) simpset = simpset addloop (name, f ctxt); |
|
147 |
fun add_solver add (ContextSolver ((name, f), _)) simpset = |
|
148 |
add (simpset, mk_solver name (f ctxt)); |
|
149 |
in |
|
150 |
((case subgoal_tac of None => ss | Some tac => ss setsubgoaler tac ctxt) |
|
151 |
addsimprocs map prep_simproc simprocs) |
|
152 |
|> fold_rev add_loop loop_tacs |
|
153 |
|> fold_rev (add_solver (op addSolver)) unsafe_solvers |
|
154 |
|> fold_rev (add_solver (op addSSolver)) solvers |
|
155 |
end; |
|
156 |
||
157 |
fun make_context_ss (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) = |
|
158 |
ContextSS {simprocs = simprocs, subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, |
|
159 |
unsafe_solvers = unsafe_solvers, solvers = solvers}; |
|
160 |
||
161 |
val empty_context_ss = make_context_ss ([], None, [], [], []); |
|
162 |
||
163 |
fun merge_context_ss (ctxt_ss1, ctxt_ss2) = |
|
164 |
let |
|
165 |
val ContextSS {simprocs = simprocs1, subgoal_tac = subgoal_tac1, loop_tacs = loop_tacs1, |
|
166 |
unsafe_solvers = unsafe_solvers1, solvers = solvers1} = ctxt_ss1; |
|
167 |
val ContextSS {simprocs = simprocs2, subgoal_tac = subgoal_tac2, loop_tacs = loop_tacs2, |
|
168 |
unsafe_solvers = unsafe_solvers2, solvers = solvers2} = ctxt_ss2; |
|
169 |
||
170 |
val simprocs' = merge_context_simprocs simprocs1 simprocs2; |
|
171 |
val subgoal_tac' = (case subgoal_tac1 of None => subgoal_tac2 | some => some); |
|
172 |
val loop_tacs' = merge_alists loop_tacs1 loop_tacs2; |
|
173 |
val unsafe_solvers' = merge_context_solvers unsafe_solvers1 unsafe_solvers2; |
|
174 |
val solvers' = merge_context_solvers solvers1 solvers2; |
|
175 |
in make_context_ss (simprocs', subgoal_tac', loop_tacs', unsafe_solvers', solvers') end; |
|
176 |
||
177 |
||
178 |
||
5024 | 179 |
(** global and local simpset data **) |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
180 |
|
5842 | 181 |
(* theory data kind 'Provers/simpset' *) |
4795 | 182 |
|
5842 | 183 |
structure GlobalSimpsetArgs = |
5024 | 184 |
struct |
5842 | 185 |
val name = "Provers/simpset"; |
15036 | 186 |
type T = simpset ref * context_ss; |
0 | 187 |
|
15036 | 188 |
val empty = (ref empty_ss, empty_context_ss); |
189 |
fun copy (ref ss, ctxt_ss) = (ref ss, ctxt_ss): T; (*create new reference!*) |
|
6556 | 190 |
val prep_ext = copy; |
15036 | 191 |
fun merge ((ref ss1, ctxt_ss1), (ref ss2, ctxt_ss2)) = |
192 |
(ref (merge_ss (ss1, ss2)), merge_context_ss (ctxt_ss1, ctxt_ss2)); |
|
193 |
fun print _ (ref ss, _) = print_ss ss; |
|
5024 | 194 |
end; |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
195 |
|
5842 | 196 |
structure GlobalSimpset = TheoryDataFun(GlobalSimpsetArgs); |
197 |
val print_simpset = GlobalSimpset.print; |
|
15036 | 198 |
val simpset_ref_of_sg = #1 o GlobalSimpset.get_sg; |
199 |
val simpset_ref_of = #1 o GlobalSimpset.get; |
|
200 |
val get_context_ss = #2 o GlobalSimpset.get o ProofContext.theory_of; |
|
201 |
||
202 |
fun map_context_ss f = GlobalSimpset.map (apsnd |
|
203 |
(fn ContextSS {simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers} => |
|
204 |
make_context_ss (f (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers)))); |
|
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
205 |
|
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
206 |
|
4795 | 207 |
(* access global simpset *) |
0 | 208 |
|
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
209 |
val simpset_of_sg = ! o simpset_ref_of_sg; |
6391 | 210 |
val simpset_of = simpset_of_sg o Theory.sign_of; |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
211 |
|
6391 | 212 |
fun SIMPSET tacf state = tacf (simpset_of_sg (Thm.sign_of_thm state)) state; |
213 |
fun SIMPSET' tacf i state = tacf (simpset_of_sg (Thm.sign_of_thm state)) i state; |
|
0 | 214 |
|
5028 | 215 |
val simpset = simpset_of o Context.the_context; |
6391 | 216 |
val simpset_ref = simpset_ref_of_sg o Theory.sign_of o Context.the_context; |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
217 |
|
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
218 |
|
4795 | 219 |
(* change global simpset *) |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
220 |
|
7177 | 221 |
fun change_simpset_of f x thy = |
222 |
let val r = simpset_ref_of thy |
|
223 |
in r := f (! r, x); thy end; |
|
224 |
||
225 |
fun change_simpset f x = (change_simpset_of f x (Context.the_context ()); ()); |
|
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
226 |
|
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
227 |
val Addsimps = change_simpset (op addsimps); |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
228 |
val Delsimps = change_simpset (op delsimps); |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
229 |
val Addsimprocs = change_simpset (op addsimprocs); |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
230 |
val Delsimprocs = change_simpset (op delsimprocs); |
9715 | 231 |
val Addcongs = change_simpset (op addcongs); |
232 |
val Delcongs = change_simpset (op delcongs); |
|
2509 | 233 |
|
0 | 234 |
|
15036 | 235 |
(* change context dependent components *) |
236 |
||
237 |
fun add_context_simprocs procs = |
|
238 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
239 |
(merge_context_simprocs procs simprocs, subgoal_tac, loop_tacs, |
|
240 |
unsafe_solvers, solvers)); |
|
241 |
||
242 |
fun del_context_simprocs procs = |
|
243 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
244 |
(gen_rems eq_context_simproc (simprocs, procs), subgoal_tac, loop_tacs, |
|
245 |
unsafe_solvers, solvers)); |
|
246 |
||
247 |
fun set_context_subgoaler tac = |
|
248 |
map_context_ss (fn (simprocs, _, loop_tacs, unsafe_solvers, solvers) => |
|
249 |
(simprocs, Some tac, loop_tacs, unsafe_solvers, solvers)); |
|
250 |
||
251 |
val reset_context_subgoaler = |
|
252 |
map_context_ss (fn (simprocs, _, loop_tacs, unsafe_solvers, solvers) => |
|
253 |
(simprocs, None, loop_tacs, unsafe_solvers, solvers)); |
|
254 |
||
255 |
fun add_context_looper (name, tac) = |
|
256 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
257 |
(simprocs, subgoal_tac, merge_alists [(name, tac)] loop_tacs, |
|
258 |
unsafe_solvers, solvers)); |
|
259 |
||
260 |
fun del_context_looper name = |
|
261 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
262 |
(simprocs, subgoal_tac, filter_out (equal name o #1) loop_tacs, |
|
263 |
unsafe_solvers, solvers)); |
|
264 |
||
265 |
fun add_context_unsafe_solver solver = |
|
266 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
267 |
(simprocs, subgoal_tac, loop_tacs, |
|
268 |
merge_context_solvers [solver] unsafe_solvers, solvers)); |
|
269 |
||
270 |
fun add_context_safe_solver solver = |
|
271 |
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) => |
|
272 |
(simprocs, subgoal_tac, loop_tacs, unsafe_solvers, |
|
273 |
merge_context_solvers [solver] solvers)); |
|
274 |
||
275 |
||
5842 | 276 |
(* proof data kind 'Provers/simpset' *) |
4795 | 277 |
|
5842 | 278 |
structure LocalSimpsetArgs = |
279 |
struct |
|
280 |
val name = "Provers/simpset"; |
|
281 |
type T = simpset; |
|
282 |
val init = simpset_of; |
|
15036 | 283 |
fun print ctxt ss = print_ss (context_ss ctxt ss (get_context_ss ctxt)); |
5842 | 284 |
end; |
4795 | 285 |
|
5842 | 286 |
structure LocalSimpset = ProofDataFun(LocalSimpsetArgs); |
287 |
val print_local_simpset = LocalSimpset.print; |
|
288 |
val get_local_simpset = LocalSimpset.get; |
|
289 |
val put_local_simpset = LocalSimpset.put; |
|
7273 | 290 |
fun map_local_simpset f ctxt = put_local_simpset (f (get_local_simpset ctxt)) ctxt; |
4795 | 291 |
|
15036 | 292 |
fun local_simpset_of ctxt = |
293 |
context_ss ctxt (get_local_simpset ctxt) (get_context_ss ctxt); |
|
294 |
||
4795 | 295 |
|
5886 | 296 |
(* attributes *) |
297 |
||
298 |
fun change_global_ss f (thy, th) = |
|
299 |
let val r = simpset_ref_of thy |
|
6096 | 300 |
in r := f (! r, [th]); (thy, th) end; |
5886 | 301 |
|
302 |
fun change_local_ss f (ctxt, th) = |
|
6096 | 303 |
let val ss = f (get_local_simpset ctxt, [th]) |
5886 | 304 |
in (put_local_simpset ss ctxt, th) end; |
305 |
||
306 |
val simp_add_global = change_global_ss (op addsimps); |
|
307 |
val simp_del_global = change_global_ss (op delsimps); |
|
308 |
val simp_add_local = change_local_ss (op addsimps); |
|
309 |
val simp_del_local = change_local_ss (op delsimps); |
|
310 |
||
9715 | 311 |
val cong_add_global = change_global_ss (op addcongs); |
312 |
val cong_del_global = change_global_ss (op delcongs); |
|
313 |
val cong_add_local = change_local_ss (op addcongs); |
|
314 |
val cong_del_local = change_local_ss (op delcongs); |
|
315 |
||
5886 | 316 |
|
10004 | 317 |
val simp_tac = generic_simp_tac false (false, false, false); |
318 |
val asm_simp_tac = generic_simp_tac false (false, true, false); |
|
319 |
val full_simp_tac = generic_simp_tac false (true, false, false); |
|
320 |
val asm_lr_simp_tac = generic_simp_tac false (true, true, false); |
|
321 |
val asm_full_simp_tac = generic_simp_tac false (true, true, true); |
|
322 |
val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); |
|
0 | 323 |
|
2629 | 324 |
|
15006
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents:
14814
diff
changeset
|
325 |
|
4795 | 326 |
(*the abstraction over the proof state delays the dereferencing*) |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
327 |
fun Simp_tac i st = simp_tac (simpset ()) i st; |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
328 |
fun Asm_simp_tac i st = asm_simp_tac (simpset ()) i st; |
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
329 |
fun Full_simp_tac i st = full_simp_tac (simpset ()) i st; |
4722
d2e44673c21e
The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents:
4713
diff
changeset
|
330 |
fun Asm_lr_simp_tac i st = asm_lr_simp_tac (simpset ()) i st; |
4080
7dce11095b0a
new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents:
3728
diff
changeset
|
331 |
fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st; |
406 | 332 |
|
5082 | 333 |
val simplify = simp_thm (false, false, false); |
334 |
val asm_simplify = simp_thm (false, true, false); |
|
335 |
val full_simplify = simp_thm (true, false, false); |
|
14814
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
336 |
val asm_lr_simplify = simp_thm (true, true, false); |
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
337 |
val asm_full_simplify = simp_thm (true, true, true); |
5082 | 338 |
|
339 |
val rewrite = simp_cterm (false, false, false); |
|
340 |
val asm_rewrite = simp_cterm (false, true, false); |
|
341 |
val full_rewrite = simp_cterm (true, false, false); |
|
14814
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
342 |
val asm_lr_rewrite = simp_cterm (true, true, false); |
c6b91c8aee1d
added asm_lr_simplify/asm_lr_rewrite and adapted asm_full_simplify/asm_full_rewrite to match to corresponding simp_tacs
schirmer
parents:
13605
diff
changeset
|
343 |
val asm_full_rewrite = simp_cterm (true, true, true); |
3557 | 344 |
|
345 |
||
4795 | 346 |
|
5886 | 347 |
(** concrete syntax of attributes **) |
5842 | 348 |
|
5886 | 349 |
(* add / del *) |
5842 | 350 |
|
5928 | 351 |
val simpN = "simp"; |
9715 | 352 |
val congN = "cong"; |
10412
1a1b4c1b2b7c
Moved meta simplification stuff from Thm to MetaSimplifier.
berghofe
parents:
10034
diff
changeset
|
353 |
val addN = "add"; |
1a1b4c1b2b7c
Moved meta simplification stuff from Thm to MetaSimplifier.
berghofe
parents:
10034
diff
changeset
|
354 |
val delN = "del"; |
6911 | 355 |
val onlyN = "only"; |
9899 | 356 |
val no_asmN = "no_asm"; |
357 |
val no_asm_useN = "no_asm_use"; |
|
358 |
val no_asm_simpN = "no_asm_simp"; |
|
13605 | 359 |
val asm_lrN = "asm_lr"; |
5842 | 360 |
|
8634 | 361 |
val simp_attr = |
362 |
(Attrib.add_del_args simp_add_global simp_del_global, |
|
363 |
Attrib.add_del_args simp_add_local simp_del_local); |
|
5842 | 364 |
|
9715 | 365 |
val cong_attr = |
366 |
(Attrib.add_del_args cong_add_global cong_del_global, |
|
367 |
Attrib.add_del_args cong_add_local cong_del_local); |
|
368 |
||
5842 | 369 |
|
370 |
(* conversions *) |
|
371 |
||
9899 | 372 |
local |
373 |
||
374 |
fun conv_mode x = |
|
375 |
((Args.parens (Args.$$$ no_asmN) >> K simplify || |
|
376 |
Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || |
|
377 |
Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || |
|
378 |
Scan.succeed asm_full_simplify) |> Scan.lift) x; |
|
379 |
||
11938 | 380 |
fun simplified_att get args = |
381 |
Attrib.syntax (conv_mode -- args >> (fn (f, ths) => |
|
382 |
Drule.rule_attribute (fn x => f ((if null ths then I else clear_ss) (get x) addsimps ths)))); |
|
9899 | 383 |
|
384 |
in |
|
385 |
||
386 |
val simplified_attr = |
|
11938 | 387 |
(simplified_att simpset_of Attrib.global_thmss, |
15036 | 388 |
simplified_att local_simpset_of Attrib.local_thmss); |
9899 | 389 |
|
390 |
end; |
|
5842 | 391 |
|
392 |
||
393 |
(* setup_attrs *) |
|
394 |
||
395 |
val setup_attrs = Attrib.add_attributes |
|
9899 | 396 |
[(simpN, simp_attr, "declaration of simplification rule"), |
397 |
(congN, cong_attr, "declaration of Simplifier congruence rule"), |
|
398 |
("simplified", simplified_attr, "simplified rule")]; |
|
5842 | 399 |
|
400 |
||
401 |
||
402 |
(** proof methods **) |
|
403 |
||
404 |
(* simplification *) |
|
405 |
||
8700 | 406 |
val simp_options = |
9899 | 407 |
(Args.parens (Args.$$$ no_asmN) >> K simp_tac || |
408 |
Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || |
|
409 |
Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || |
|
13605 | 410 |
Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || |
8700 | 411 |
Scan.succeed asm_full_simp_tac); |
412 |
||
9715 | 413 |
val cong_modifiers = |
9720 | 414 |
[Args.$$$ congN -- Args.colon >> K ((I, cong_add_local):Method.modifier), |
10034 | 415 |
Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add_local), |
416 |
Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del_local)]; |
|
9715 | 417 |
|
5928 | 418 |
val simp_modifiers = |
8815 | 419 |
[Args.$$$ simpN -- Args.colon >> K (I, simp_add_local), |
10034 | 420 |
Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add_local), |
421 |
Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del_local), |
|
9861 | 422 |
Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> K (map_local_simpset clear_ss, simp_add_local)] |
423 |
@ cong_modifiers; |
|
5928 | 424 |
|
425 |
val simp_modifiers' = |
|
10034 | 426 |
[Args.add -- Args.colon >> K (I, simp_add_local), |
427 |
Args.del -- Args.colon >> K (I, simp_del_local), |
|
9861 | 428 |
Args.$$$ onlyN -- Args.colon >> K (map_local_simpset clear_ss, simp_add_local)] |
429 |
@ cong_modifiers; |
|
5928 | 430 |
|
8700 | 431 |
fun simp_args more_mods = |
432 |
Method.sectioned_args (Args.bang_facts -- Scan.lift simp_options) (more_mods @ simp_modifiers'); |
|
433 |
||
8170 | 434 |
|
8700 | 435 |
fun simp_method (prems, tac) ctxt = Method.METHOD (fn facts => |
436 |
ALLGOALS (Method.insert_tac (prems @ facts)) THEN |
|
15036 | 437 |
(CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt)); |
8700 | 438 |
|
439 |
fun simp_method' (prems, tac) ctxt = Method.METHOD (fn facts => |
|
10821 | 440 |
HEADGOAL (Method.insert_tac (prems @ facts) THEN' |
15036 | 441 |
(CHANGED_PROP oo tac) (local_simpset_of ctxt))); |
9715 | 442 |
|
5842 | 443 |
|
444 |
(* setup_methods *) |
|
445 |
||
8467 | 446 |
fun setup_methods more_mods = Method.add_methods |
8700 | 447 |
[(simpN, simp_args more_mods simp_method', "simplification"), |
448 |
("simp_all", simp_args more_mods simp_method, "simplification (all goals)")]; |
|
5842 | 449 |
|
450 |
||
451 |
||
4795 | 452 |
(** theory setup **) |
453 |
||
8467 | 454 |
val setup = [GlobalSimpset.init, LocalSimpset.init, setup_attrs]; |
455 |
fun method_setup mods = [setup_methods mods]; |
|
4795 | 456 |
|
8228 | 457 |
fun easy_setup reflect trivs = |
458 |
let |
|
459 |
val trivialities = Drule.reflexive_thm :: trivs; |
|
460 |
||
461 |
fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac]; |
|
462 |
val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; |
|
463 |
||
464 |
(*no premature instantiation of variables during simplification*) |
|
465 |
fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac]; |
|
466 |
val safe_solver = mk_solver "easy safe" safe_solver_tac; |
|
467 |
||
8243 | 468 |
fun mk_eq thm = |
8228 | 469 |
if Logic.is_equals (Thm.concl_of thm) then [thm] |
470 |
else [thm RS reflect] handle THM _ => []; |
|
471 |
||
8243 | 472 |
fun mksimps thm = mk_eq (Drule.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm); |
473 |
||
8228 | 474 |
fun init_ss thy = |
475 |
(simpset_ref_of thy := |
|
476 |
empty_ss setsubgoaler asm_simp_tac |
|
477 |
setSSolver safe_solver |
|
478 |
setSolver unsafe_solver |
|
479 |
setmksimps mksimps; thy); |
|
8467 | 480 |
in setup @ method_setup [] @ [init_ss] end; |
8228 | 481 |
|
4795 | 482 |
|
8667 | 483 |
|
484 |
(** outer syntax **) |
|
485 |
||
486 |
val print_simpsetP = |
|
487 |
OuterSyntax.improper_command "print_simpset" "print context of Simplifier" |
|
488 |
OuterSyntax.Keyword.diag |
|
9513 | 489 |
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o (Toplevel.keep |
9010 | 490 |
(Toplevel.node_case print_simpset (print_local_simpset o Proof.context_of))))); |
8667 | 491 |
|
492 |
val _ = OuterSyntax.add_parsers [print_simpsetP]; |
|
493 |
||
1243 | 494 |
end; |
4795 | 495 |
|
496 |
||
497 |
structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier; |
|
498 |
open BasicSimplifier; |