1 (* Title: Pure/simplifier.ML
3 Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
5 Generic simplifier, suitable for most logics (see also
6 meta_simplifier.ML for the actual meta-level rewriting engine).
9 signature BASIC_SIMPLIFIER =
11 include BASIC_META_SIMPLIFIER
12 val print_simpset: theory -> unit
13 val change_simpset_of: theory -> (simpset -> simpset) -> unit
14 val change_simpset: (simpset -> simpset) -> unit
15 val simpset_of: theory -> simpset
16 val simpset: unit -> simpset
17 val SIMPSET: (simpset -> tactic) -> tactic
18 val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic
19 val Addsimps: thm list -> unit
20 val Delsimps: thm list -> unit
21 val Addsimprocs: simproc list -> unit
22 val Delsimprocs: simproc list -> unit
23 val Addcongs: thm list -> unit
24 val Delcongs: thm list -> unit
25 val local_simpset_of: Proof.context -> simpset
26 val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
27 val safe_asm_full_simp_tac: simpset -> int -> tactic
28 val simp_tac: simpset -> int -> tactic
29 val asm_simp_tac: simpset -> int -> tactic
30 val full_simp_tac: simpset -> int -> tactic
31 val asm_lr_simp_tac: simpset -> int -> tactic
32 val asm_full_simp_tac: simpset -> int -> tactic
33 val Simp_tac: int -> tactic
34 val Asm_simp_tac: int -> tactic
35 val Full_simp_tac: int -> tactic
36 val Asm_lr_simp_tac: int -> tactic
37 val Asm_full_simp_tac: int -> tactic
38 val simplify: simpset -> thm -> thm
39 val asm_simplify: simpset -> thm -> thm
40 val full_simplify: simpset -> thm -> thm
41 val asm_lr_simplify: simpset -> thm -> thm
42 val asm_full_simplify: simpset -> thm -> thm
45 signature SIMPLIFIER =
47 include BASIC_SIMPLIFIER
48 val clear_ss: simpset -> simpset
49 val debug_bounds: bool ref
50 val inherit_context: simpset -> simpset -> simpset
51 val the_context: simpset -> Proof.context
52 val context: Proof.context -> simpset -> simpset
53 val theory_context: theory -> simpset -> simpset
54 val simproc_i: theory -> string -> term list
55 -> (theory -> simpset -> term -> thm option) -> simproc
56 val simproc: theory -> string -> string list
57 -> (theory -> simpset -> term -> thm option) -> simproc
58 val rewrite: simpset -> cterm -> thm
59 val asm_rewrite: simpset -> cterm -> thm
60 val full_rewrite: simpset -> cterm -> thm
61 val asm_lr_rewrite: simpset -> cterm -> thm
62 val asm_full_rewrite: simpset -> cterm -> thm
63 val get_simpset: theory -> simpset
64 val print_local_simpset: Proof.context -> unit
65 val get_local_simpset: Proof.context -> simpset
66 val put_local_simpset: simpset -> Proof.context -> Proof.context
67 val attrib: (simpset * thm list -> simpset) -> attribute
68 val simp_add: attribute
69 val simp_del: attribute
70 val cong_add: attribute
71 val cong_del: attribute
72 val get_simproc: Proof.context -> xstring -> simproc
73 val def_simproc: string -> string list -> (morphism -> simpset -> cterm -> thm option) ->
74 local_theory -> local_theory
75 val def_simproc_i: string -> term list -> (morphism -> simpset -> cterm -> thm option) ->
76 local_theory -> local_theory
77 val cong_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
78 val simp_modifiers': (Args.T list -> (Method.modifier * Args.T list)) list
79 val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
80 val method_setup: (Args.T list -> (Method.modifier * Args.T list)) list
82 val easy_setup: thm -> thm list -> theory -> theory
85 structure Simplifier: SIMPLIFIER =
95 structure GlobalSimpset = TheoryDataFun
97 val name = "Pure/simpset";
100 val empty = ref empty_ss;
101 fun copy (ref ss) = ref ss: T; (*create new reference!*)
102 fun extend (ref ss) = ref (MetaSimplifier.inherit_context empty_ss ss);
103 fun merge _ (ref ss1, ref ss2) = ref (merge_ss (ss1, ss2));
104 fun print _ (ref ss) = print_ss ss;
107 val _ = Context.add_setup GlobalSimpset.init;
108 val print_simpset = GlobalSimpset.print;
109 val get_simpset = ! o GlobalSimpset.get;
111 val change_simpset_of = change o GlobalSimpset.get;
112 fun change_simpset f = change_simpset_of (ML_Context.the_context ()) f;
114 fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_simpset thy);
115 val simpset = simpset_of o ML_Context.the_context;
118 fun SIMPSET tacf st = tacf (simpset_of (Thm.theory_of_thm st)) st;
119 fun SIMPSET' tacf i st = tacf (simpset_of (Thm.theory_of_thm st)) i st;
121 fun Addsimps args = change_simpset (fn ss => ss addsimps args);
122 fun Delsimps args = change_simpset (fn ss => ss delsimps args);
123 fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args);
124 fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args);
125 fun Addcongs args = change_simpset (fn ss => ss addcongs args);
126 fun Delcongs args = change_simpset (fn ss => ss delcongs args);
131 structure LocalSimpset = ProofDataFun
133 val name = "Pure/simpset";
135 val init = get_simpset;
136 fun print _ ss = print_ss ss;
139 val _ = Context.add_setup LocalSimpset.init;
140 val print_local_simpset = LocalSimpset.print;
141 val get_local_simpset = LocalSimpset.get;
142 val put_local_simpset = LocalSimpset.put;
144 fun local_simpset_of ctxt = MetaSimplifier.context ctxt (get_local_simpset ctxt);
146 val _ = ML_Context.value_antiq "simpset"
147 (Scan.succeed ("simpset", "Simplifier.local_simpset_of (ML_Context.the_local_context ())"));
150 (* generic simpsets *)
152 fun get_ss (Context.Theory thy) = simpset_of thy
153 | get_ss (Context.Proof ctxt) = local_simpset_of ctxt;
155 fun map_ss f (Context.Theory thy) = (change_simpset_of thy f; Context.Theory thy)
156 | map_ss f (Context.Proof ctxt) = Context.Proof (LocalSimpset.map f ctxt);
161 fun attrib f = Thm.declaration_attribute (fn th => map_ss (fn ss => f (ss, [th])));
163 val simp_add = attrib (op addsimps);
164 val simp_del = attrib (op delsimps);
165 val cong_add = attrib (op addcongs);
166 val cong_del = attrib (op delcongs);
170 (** named simprocs **)
172 fun err_dup_simprocs ds =
173 error ("Duplicate simproc(s): " ^ commas_quote ds);
177 structure Simprocs = GenericDataFun
179 val name = "Pure/simprocs";
180 type T = (simproc * stamp) NameSpace.table;
181 val empty = NameSpace.empty_table;
183 fun merge _ simprocs = NameSpace.merge_tables (eq_snd (op =)) simprocs
184 handle Symtab.DUPS ds => err_dup_simprocs ds;
187 val _ = Context.add_setup Simprocs.init;
192 fun get_simproc ctxt xname =
194 val (space, tab) = Simprocs.get (Context.Proof ctxt);
195 val name = NameSpace.intern space xname;
197 (case Symtab.lookup tab name of
198 SOME (proc, _) => proc
199 | NONE => error ("Undefined simplification procedure: " ^ quote name))
202 val _ = ML_Context.value_antiq "simproc" (Scan.lift Args.name >> (fn name =>
204 "Simplifier.get_simproc (ML_Context.the_local_context ()) " ^ ML_Syntax.print_string name)));
207 (* define simprocs *)
212 fun read_terms ctxt ts =
213 #1 (ProofContext.read_termTs ctxt (K false) (K NONE) (K NONE) []
214 (map (rpair TypeInfer.logicT) ts));
217 fun cert_terms ctxt ts = map (ProofContext.cert_term ctxt) ts;
219 fun prep_pats prep ctxt raw_pats =
221 val pats = prep ctxt raw_pats;
223 |> fold Variable.declare_term pats
224 |> fold Variable.auto_fixes pats;
225 in Variable.export_terms ctxt' ctxt pats end;
227 fun gen_simproc prep name raw_pats proc lthy =
230 prep_pats prep lthy raw_pats
231 |> map (Morphism.term (LocalTheory.target_morphism lthy));
232 val naming = LocalTheory.target_naming lthy;
234 lthy |> LocalTheory.declaration (fn phi => fn context =>
236 val cert = Thm.cterm_of (Context.theory_of context);
237 val pats' = map (cert o Morphism.term phi) pats;
238 val name' = Morphism.name phi name;
239 val proc' = proc phi;
240 val simproc = mk_simproc' (NameSpace.full naming name') pats' proc';
243 |> Simprocs.map (fn simprocs =>
244 NameSpace.extend_table naming (simprocs, [(name', (simproc, stamp ()))])
245 handle Symtab.DUPS ds => err_dup_simprocs ds)
246 |> map_ss (fn ss => ss addsimprocs [simproc])
252 val def_simproc = gen_simproc read_terms;
253 val def_simproc_i = gen_simproc cert_terms;
259 (** simplification tactics and rules **)
261 fun solve_all_tac solvers ss =
263 val (_, {subgoal_tac, ...}) = MetaSimplifier.rep_ss ss;
264 val solve_tac = subgoal_tac (MetaSimplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac);
265 in DEPTH_SOLVE (solve_tac 1) end;
267 (*NOTE: may instantiate unknowns that appear also in other subgoals*)
268 fun generic_simp_tac safe mode ss =
270 val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = MetaSimplifier.rep_ss ss;
271 val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs));
272 val solve_tac = FIRST' (map (MetaSimplifier.solver ss)
273 (if safe then solvers else unsafe_solvers));
275 fun simp_loop_tac i =
276 Goal.asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN
277 (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));
278 in simp_loop_tac end;
282 fun simp rew mode ss thm =
284 val (_, {solvers = (unsafe_solvers, _), ...}) = MetaSimplifier.rep_ss ss;
285 val tacf = solve_all_tac unsafe_solvers;
286 fun prover s th = Option.map #1 (Seq.pull (tacf s th));
287 in rew mode prover ss thm end;
291 val simp_thm = simp MetaSimplifier.rewrite_thm;
292 val simp_cterm = simp MetaSimplifier.rewrite_cterm;
299 val simp_tac = generic_simp_tac false (false, false, false);
300 val asm_simp_tac = generic_simp_tac false (false, true, false);
301 val full_simp_tac = generic_simp_tac false (true, false, false);
302 val asm_lr_simp_tac = generic_simp_tac false (true, true, false);
303 val asm_full_simp_tac = generic_simp_tac false (true, true, true);
304 val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);
306 (*the abstraction over the proof state delays the dereferencing*)
307 fun Simp_tac i st = simp_tac (simpset ()) i st;
308 fun Asm_simp_tac i st = asm_simp_tac (simpset ()) i st;
309 fun Full_simp_tac i st = full_simp_tac (simpset ()) i st;
310 fun Asm_lr_simp_tac i st = asm_lr_simp_tac (simpset ()) i st;
311 fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st;
316 val simplify = simp_thm (false, false, false);
317 val asm_simplify = simp_thm (false, true, false);
318 val full_simplify = simp_thm (true, false, false);
319 val asm_lr_simplify = simp_thm (true, true, false);
320 val asm_full_simplify = simp_thm (true, true, true);
322 val rewrite = simp_cterm (false, false, false);
323 val asm_rewrite = simp_cterm (false, true, false);
324 val full_rewrite = simp_cterm (true, false, false);
325 val asm_lr_rewrite = simp_cterm (true, true, false);
326 val asm_full_rewrite = simp_cterm (true, true, true);
330 (** concrete syntax of attributes **)
339 val no_asmN = "no_asm";
340 val no_asm_useN = "no_asm_use";
341 val no_asm_simpN = "no_asm_simp";
342 val asm_lrN = "asm_lr";
350 ((Args.parens (Args.$$$ no_asmN) >> K simplify ||
351 Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify ||
352 Args.parens (Args.$$$ no_asm_useN) >> K full_simplify ||
353 Scan.succeed asm_full_simplify) |> Scan.lift) x;
358 Attrib.syntax (conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute (fn x =>
359 f ((if null ths then I else MetaSimplifier.clear_ss) (get_ss x) addsimps ths))));
364 (* setup attributes *)
366 val _ = Context.add_setup
367 (Attrib.add_attributes
368 [(simpN, Attrib.add_del_args simp_add simp_del, "declaration of Simplifier rule"),
369 (congN, Attrib.add_del_args cong_add cong_del, "declaration of Simplifier congruence rule"),
370 ("simplified", simplified, "simplified rule")]);
374 (** proof methods **)
379 (Args.parens (Args.$$$ no_asmN) >> K simp_tac ||
380 Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac ||
381 Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac ||
382 Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac ||
383 Scan.succeed asm_full_simp_tac);
385 fun simp_flags x = (Scan.repeat
386 (Args.parens (Args.$$$ "depth_limit" -- Args.colon |-- Args.nat)
387 >> setmp MetaSimplifier.simp_depth_limit)
388 >> (curry (Library.foldl op o) I o rev)) x;
391 [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),
392 Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),
393 Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];
396 [Args.$$$ simpN -- Args.colon >> K (I, simp_add),
397 Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),
398 Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),
399 Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon
400 >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add)]
403 val simp_modifiers' =
404 [Args.add -- Args.colon >> K (I, simp_add),
405 Args.del -- Args.colon >> K (I, simp_del),
406 Args.$$$ onlyN -- Args.colon
407 >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add)]
410 fun simp_args more_mods =
411 Method.sectioned_args (Args.bang_facts -- Scan.lift simp_options -- Scan.lift simp_flags)
412 (more_mods @ simp_modifiers');
414 fun simp_method ((prems, tac), FLAGS) ctxt = Method.METHOD (fn facts =>
415 ALLGOALS (Method.insert_tac (prems @ facts)) THEN
416 (FLAGS o CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt));
418 fun simp_method' ((prems, tac), FLAGS) ctxt = Method.METHOD (fn facts =>
419 HEADGOAL (Method.insert_tac (prems @ facts) THEN'
420 ((FLAGS o CHANGED_PROP) oo tac) (local_simpset_of ctxt)));
426 fun method_setup more_mods = Method.add_methods
427 [(simpN, simp_args more_mods simp_method', "simplification"),
428 ("simp_all", simp_args more_mods simp_method, "simplification (all goals)")];
430 fun easy_setup reflect trivs = method_setup [] #> (fn thy =>
432 val trivialities = Drule.reflexive_thm :: trivs;
434 fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac];
435 val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;
437 (*no premature instantiation of variables during simplification*)
438 fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac];
439 val safe_solver = mk_solver "easy safe" safe_solver_tac;
442 if can Logic.dest_equals (Thm.concl_of thm) then [thm]
443 else [thm RS reflect] handle THM _ => [];
445 fun mksimps thm = mk_eq (Drule.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm);
447 GlobalSimpset.get thy :=
448 empty_ss setsubgoaler asm_simp_tac
449 setSSolver safe_solver
450 setSolver unsafe_solver
457 structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier;
458 open BasicSimplifier;