1 (* Title: Pure/meta_simplifier.ML
3 Author: Tobias Nipkow and Stefan Berghofer
5 Meta-level Simplification.
9 addsimps delsimps addeqcongs deleqcongs addcongs delcongs addsimprocs delsimprocs
10 setmksimps setmksimps2 setmkcong setmksym setmkeqTrue settermless setsubgoaler
11 setloop addloop delloop setSSolver addSSolver setSolver addSolver;
13 signature BASIC_META_SIMPLIFIER =
15 val debug_simp: bool ref
16 val trace_simp: bool ref
17 val simp_depth_limit: int ref
21 val mk_solver: string -> (thm list -> int -> tactic) -> solver
24 val rep_ss: simpset ->
25 {rules: rrule Net.net,
29 {congs: (string * cong) list * string list,
33 mk2: thm list -> thm -> thm list,
35 mk_sym: thm -> thm option,
36 mk_eq_True: thm -> thm option},
37 termless: term * term -> bool,
38 subgoal_tac: simpset -> int -> tactic,
39 loop_tacs: (string * (int -> tactic)) list,
40 solvers: solver list * solver list}
41 val print_ss: simpset -> unit
43 val merge_ss: simpset * simpset -> simpset
45 val mk_simproc: string -> cterm list ->
46 (Sign.sg -> simpset -> term -> thm option) -> simproc
47 val add_prems: thm list -> simpset -> simpset
48 val prems_of_ss: simpset -> thm list
49 val addsimps: simpset * thm list -> simpset
50 val delsimps: simpset * thm list -> simpset
51 val addeqcongs: simpset * thm list -> simpset
52 val deleqcongs: simpset * thm list -> simpset
53 val addcongs: simpset * thm list -> simpset
54 val delcongs: simpset * thm list -> simpset
55 val addsimprocs: simpset * simproc list -> simpset
56 val delsimprocs: simpset * simproc list -> simpset
57 val setmksimps: simpset * (thm -> thm list) -> simpset
58 val setmksimps2: simpset * (thm list -> thm -> thm list) -> simpset
59 val setmkcong: simpset * (thm -> thm) -> simpset
60 val setmksym: simpset * (thm -> thm option) -> simpset
61 val setmkeqTrue: simpset * (thm -> thm option) -> simpset
62 val settermless: simpset * (term * term -> bool) -> simpset
63 val setsubgoaler: simpset * (simpset -> int -> tactic) -> simpset
64 val setloop: simpset * (int -> tactic) -> simpset
65 val addloop: simpset * (string * (int -> tactic)) -> simpset
66 val delloop: simpset * string -> simpset
67 val setSSolver: simpset * solver -> simpset
68 val addSSolver: simpset * solver -> simpset
69 val setSolver: simpset * solver -> simpset
70 val addSolver: simpset * solver -> simpset
71 val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
74 signature META_SIMPLIFIER =
76 include BASIC_META_SIMPLIFIER
77 exception SIMPLIFIER of string * thm
78 val clear_ss: simpset -> simpset
79 exception SIMPROC_FAIL of string * exn
80 val simproc_i: Sign.sg -> string -> term list
81 -> (Sign.sg -> simpset -> term -> thm option) -> simproc
82 val simproc: Sign.sg -> string -> string list
83 -> (Sign.sg -> simpset -> term -> thm option) -> simproc
84 val rewrite_cterm: bool * bool * bool ->
85 (simpset -> thm -> thm option) -> simpset -> cterm -> thm
86 val rewrite_aux: (simpset -> thm -> thm option) -> bool -> thm list -> cterm -> thm
87 val simplify_aux: (simpset -> thm -> thm option) -> bool -> thm list -> thm -> thm
88 val rewrite_term: Sign.sg -> thm list -> (term -> term option) list -> term -> term
89 val rewrite_thm: bool * bool * bool ->
90 (simpset -> thm -> thm option) -> simpset -> thm -> thm
91 val rewrite_goals_rule_aux: (simpset -> thm -> thm option) -> thm list -> thm -> thm
92 val rewrite_goal_rule: bool * bool * bool ->
93 (simpset -> thm -> thm option) -> simpset -> int -> thm -> thm
94 val asm_rewrite_goal_tac: bool * bool * bool ->
95 (simpset -> tactic) -> simpset -> int -> tactic
96 val simp_thm: bool * bool * bool -> simpset -> thm -> thm
97 val simp_cterm: bool * bool * bool -> simpset -> cterm -> thm
100 structure MetaSimplifier: META_SIMPLIFIER =
106 exception SIMPLIFIER of string * thm;
108 val debug_simp = ref false;
109 val trace_simp = ref false;
110 val simp_depth = ref 0;
111 val simp_depth_limit = ref 1000;
116 tracing (case ! simp_depth of 0 => a | n => enclose "[" "]" (string_of_int n) ^ a);
118 fun prnt warn a = if warn then warning a else println a;
119 fun prtm warn a sg t = prnt warn (a ^ "\n" ^ Sign.string_of_term sg t);
120 fun prctm warn a t = prnt warn (a ^ "\n" ^ Display.string_of_cterm t);
124 fun debug warn a = if ! debug_simp then prnt warn a else ();
125 fun trace warn a = if ! trace_simp then prnt warn a else ();
127 fun debug_term warn a sign t = if ! debug_simp then prtm warn a sign t else ();
128 fun trace_term warn a sign t = if ! trace_simp then prtm warn a sign t else ();
129 fun trace_cterm warn a ct = if ! trace_simp then prctm warn a ct else ();
130 fun trace_thm a th = if ! trace_simp then prctm false a (Thm.cprop_of th) else ();
132 fun trace_named_thm a (thm, name) =
134 prctm false (if name = "" then a else a ^ " " ^ quote name ^ ":") (Thm.cprop_of thm)
137 fun warn_thm a = prctm true a o Thm.cprop_of;
143 (** datatype simpset **)
147 type rrule = {thm: thm, name: string, lhs: term, elhs: cterm, fo: bool, perm: bool};
149 (*thm: the rewrite rule;
150 name: name of theorem from which rewrite rule was extracted;
151 lhs: the left-hand side;
152 elhs: the etac-contracted lhs;
153 fo: use first-order matching;
154 perm: the rewrite rule is permutative;
157 - elhs is used for matching,
158 lhs only for preservation of bound variable names;
160 either elhs is first-order (no Var is applied),
161 in which case fo-matching is complete,
162 or elhs is not a pattern,
163 in which case there is nothing better to do;*)
165 fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) =
166 Drule.eq_thm_prop (thm1, thm2);
171 type cong = {thm: thm, lhs: cterm};
173 fun eq_cong ({thm = thm1, ...}: cong, {thm = thm2, ...}: cong) =
174 Drule.eq_thm_prop (thm1, thm2);
182 solver: thm list -> int -> tactic,
185 fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()};
187 fun solver_name (Solver {name, ...}) = name;
188 fun solver ths (Solver {solver = tacf, ...}) = tacf ths;
189 fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2);
190 val merge_solvers = gen_merge_lists eq_solver;
193 (* simplification sets and procedures *)
195 (*A simpset contains data required during conversion:
196 rules: discrimination net of rewrite rules;
197 prems: current premises;
198 bounds: names of bound variables already used
199 (for generating new names when rewriting under lambda abstractions);
200 depth: depth of conditional rewriting;
201 congs: association list of congruence rules and
202 a list of `weak' congruence constants.
203 A congruence is `weak' if it avoids normalization of some argument.
204 procs: discrimination net of simplification procedures
205 (functions that prove rewrite rules on the fly);
207 mk: turn simplification thms into rewrite rules;
208 mk2: like mk but may also depend on the other premises
209 mk_cong: prepare congruence rules;
210 mk_sym: turn == around;
211 mk_eq_True: turn P into P == True;
212 termless: relation for ordered rewriting;*)
215 {mk: thm -> thm list,
216 mk2: thm list -> thm -> thm list,
218 mk_sym: thm -> thm option,
219 mk_eq_True: thm -> thm option};
223 {rules: rrule Net.net,
227 {congs: (string * cong) list * string list,
230 termless: term * term -> bool,
231 subgoal_tac: simpset -> int -> tactic,
232 loop_tacs: (string * (int -> tactic)) list,
233 solvers: solver list * solver list}
238 proc: Sign.sg -> simpset -> term -> thm option,
241 fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = (id1 = id2);
243 fun rep_ss (Simpset args) = args;
245 fun make_ss1 (rules, prems, bounds, depth) =
246 {rules = rules, prems = prems, bounds = bounds, depth = depth};
248 fun map_ss1 f {rules, prems, bounds, depth} =
249 make_ss1 (f (rules, prems, bounds, depth));
251 fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =
252 {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless,
253 subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers};
255 fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} =
256 make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
258 fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2);
260 fun map_simpset f (Simpset ({rules, prems, bounds, depth},
261 {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers})) =
262 make_simpset (f ((rules, prems, bounds, depth),
263 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)));
265 fun map_simpset1 f (Simpset (r1, r2)) = Simpset (map_ss1 f r1, r2);
266 fun map_simpset2 f (Simpset (r1, r2)) = Simpset (r1, map_ss2 f r2);
273 val pretty_thms = map Display.pretty_thm;
275 fun pretty_cong (name, th) =
276 Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, Display.pretty_thm th];
277 fun pretty_proc (name, lhss) =
278 Pretty.big_list (name ^ ":") (map Display.pretty_cterm lhss);
280 val Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...}) = ss;
281 val smps = map (#thm o #2) (Net.dest rules);
282 val cngs = map (fn (name, {thm, ...}) => (name, thm)) (#1 congs);
283 val prcs = Net.dest procs |>
284 map (fn (_, Proc {name, lhs, id, ...}) => ((name, lhs), id))
285 |> partition_eq eq_snd
286 |> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps))
287 |> Library.sort_wrt #1;
289 [Pretty.big_list "simplification rules:" (pretty_thms smps),
290 Pretty.big_list "simplification procedures:" (map pretty_proc prcs),
291 Pretty.big_list "congruences:" (map pretty_cong cngs),
292 Pretty.strs ("loopers:" :: map (quote o #1) loop_tacs),
293 Pretty.strs ("unsafe solvers:" :: map (quote o solver_name) (#1 solvers)),
294 Pretty.strs ("safe solvers:" :: map (quote o solver_name) (#2 solvers))]
295 |> Pretty.chunks |> Pretty.writeln
303 fun init_ss mk_rews termless subgoal_tac solvers =
304 make_simpset ((Net.empty, [], [], 0),
305 (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers));
307 val basic_mk_rews: mk_rews =
308 {mk = fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [],
309 mk2 = fn thms => fn thm => [],
311 mk_sym = Some o Drule.symmetric_fun,
312 mk_eq_True = K None};
316 val empty_ss = init_ss basic_mk_rews Term.termless (K (K no_tac)) ([], []);
318 fun clear_ss (Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...})) =
319 init_ss mk_rews termless subgoal_tac solvers;
324 (* merge simpsets *) (*NOTE: ignores some fields of 2nd simpset*)
326 fun merge_ss (ss1, ss2) =
328 val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth},
329 {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac,
330 loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1;
331 val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = _},
332 {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _,
333 loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2;
335 val rules' = Net.merge (rules1, rules2, eq_rrule);
336 val prems' = gen_merge_lists Drule.eq_thm_prop prems1 prems2;
337 val bounds' = merge_lists bounds1 bounds2;
338 val congs' = gen_merge_lists (eq_cong o pairself #2) congs1 congs2;
339 val weak' = merge_lists weak1 weak2;
340 val procs' = Net.merge (procs1, procs2, eq_proc);
341 val loop_tacs' = merge_alists loop_tacs1 loop_tacs2;
342 val unsafe_solvers' = merge_solvers unsafe_solvers1 unsafe_solvers2;
343 val solvers' = merge_solvers solvers1 solvers2;
345 make_simpset ((rules', prems', bounds', depth), ((congs', weak'), procs',
346 mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers')))
352 exception SIMPROC_FAIL of string * exn;
354 datatype simproc = Simproc of proc list;
356 fun mk_simproc name lhss proc =
357 let val id = stamp () in
358 Simproc (lhss |> map (fn lhs =>
359 Proc {name = name, lhs = lhs, proc = proc, id = id}))
362 fun simproc_i sg name = mk_simproc name o map (Thm.cterm_of sg o Logic.varify);
363 fun simproc sg name = simproc_i sg name o map (Sign.simple_read_term sg TypeInfer.logicT);
367 (** simpset operations **)
369 (* bounds and prems *)
371 fun add_bound b = map_simpset1 (fn (rules, prems, bounds, depth) =>
372 (rules, prems, b :: bounds, depth));
374 fun add_prems ths = map_simpset1 (fn (rules, prems, bounds, depth) =>
375 (rules, ths @ prems, bounds, depth));
377 fun prems_of_ss (Simpset ({prems, ...}, _)) = prems;
382 fun mk_rrule2 {thm, name, lhs, elhs, perm} =
384 val fo = Pattern.first_order (term_of elhs) orelse not (Pattern.pattern (term_of elhs))
385 in {thm = thm, name = name, lhs = lhs, elhs = elhs, fo = fo, perm = perm} end;
387 fun insert_rrule quiet (ss, rrule as {thm, name, lhs, elhs, perm}) =
388 (trace_named_thm "Adding rewrite rule" (thm, name);
389 ss |> map_simpset1 (fn (rules, prems, bounds, depth) =>
391 val rrule2 as {elhs, ...} = mk_rrule2 rrule;
392 val rules' = Net.insert_term ((term_of elhs, rrule2), rules, eq_rrule);
393 in (rules', prems, bounds, depth) end)
395 (if quiet then () else warn_thm "Ignoring duplicate rewrite rule:" thm; ss));
397 fun vperm (Var _, Var _) = true
398 | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t)
399 | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2)
400 | vperm (t, u) = (t = u);
402 fun var_perm (t, u) =
403 vperm (t, u) andalso eq_set (term_varnames t, term_varnames u);
405 (* FIXME: it seems that the conditions on extra variables are too liberal if
406 prems are nonempty: does solving the prems really guarantee instantiation of
407 all its Vars? Better: a dynamic check each time a rule is applied.
409 fun rewrite_rule_extra_vars prems elhs erhs =
410 not (term_varnames erhs subset foldl add_term_varnames (term_varnames elhs, prems))
412 not (term_tvars erhs subset (term_tvars elhs union List.concat (map term_tvars prems)));
414 (*simple test for looping rewrite rules and stupid orientations*)
415 fun reorient sign prems lhs rhs =
416 rewrite_rule_extra_vars prems lhs rhs
420 exists (apl (lhs, Logic.occs)) (rhs :: prems)
422 null prems andalso Pattern.matches (Sign.tsig_of sign) (lhs, rhs)
423 (*the condition "null prems" is necessary because conditional rewrites
424 with extra variables in the conditions may terminate although
425 the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*)
427 is_Const lhs andalso not (is_Const rhs);
429 fun decomp_simp thm =
431 val {sign, prop, ...} = Thm.rep_thm thm;
432 val prems = Logic.strip_imp_prems prop;
433 val concl = Drule.strip_imp_concl (Thm.cprop_of thm);
434 val (lhs, rhs) = Drule.dest_equals concl handle TERM _ =>
435 raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm);
436 val (_, elhs) = Drule.dest_equals (Thm.cprop_of (Thm.eta_conversion lhs));
437 val elhs = if elhs = lhs then lhs else elhs; (*share identical copies*)
438 val erhs = Pattern.eta_contract (term_of rhs);
440 var_perm (term_of elhs, erhs) andalso
441 not (term_of elhs aconv erhs) andalso
442 not (is_Var (term_of elhs));
443 in (sign, prems, term_of lhs, elhs, term_of rhs, perm) end;
445 fun decomp_simp' thm =
446 let val (_, _, lhs, _, rhs, _) = decomp_simp thm in
447 if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", thm)
451 fun mk_eq_True (Simpset (_, {mk_rews = {mk_eq_True, ...}, ...})) (thm, name) =
452 (case mk_eq_True thm of
455 let val (_, _, lhs, elhs, _, _) = decomp_simp eq_True
456 in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end);
458 (*create the rewrite rule and possibly also the eq_True variant,
459 in case there are extra vars on the rhs*)
460 fun rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm2) =
461 let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in
462 if term_varnames rhs subset term_varnames lhs andalso
463 term_tvars rhs subset term_tvars lhs then [rrule]
464 else mk_eq_True ss (thm2, name) @ [rrule]
467 fun mk_rrule ss (thm, name) =
468 let val (_, prems, lhs, elhs, rhs, perm) = decomp_simp thm in
469 if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
471 (*weak test for loops*)
472 if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs)
473 then mk_eq_True ss (thm, name)
474 else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm)
477 fun orient_rrule ss (thm, name) =
478 let val (sign, prems, lhs, elhs, rhs, perm) = decomp_simp thm in
479 if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
480 else if reorient sign prems lhs rhs then
481 if reorient sign prems rhs lhs
482 then mk_eq_True ss (thm, name)
484 let val Simpset (_, {mk_rews = {mk_sym, ...}, ...}) = ss in
488 let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm'
489 in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end)
491 else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm)
494 fun extract_rews (Simpset ({prems, ...}, {mk_rews = {mk, ...}, ...}), thms)
495 = flat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms);
497 fun extract_rews2 (Simpset ({prems, ...}, {mk_rews = {mk2, ...}, ...}), thms)
498 = flat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk2 prems thm)) thms);
500 fun orient_comb_simps comb mk_rrule (ss, thms) =
502 val rews = extract_rews (ss, thms);
503 val rrules = flat (map mk_rrule rews);
504 in foldl comb (ss, rrules) end;
506 fun extract_safe_rrules (ss, thm) =
507 flat (map (orient_rrule ss) (extract_rews (ss, [thm])));
509 fun extract_safe_rrules2 (ss, thm) =
510 flat (map (orient_rrule ss) (extract_rews2 (ss, [thm])));
512 (*add rewrite rules explicitly; do not reorient!*)
513 fun ss addsimps thms =
514 orient_comb_simps (insert_rrule false) (mk_rrule ss) (ss, thms);
517 fun add_prem2(ss,thm) =
518 foldl (insert_rrule true) (ss,extract_safe_rrules2(ss,thm))
521 fun add_prems2 thms ss = foldl add_prem2 (ss,thms);
525 fun del_rrule (ss, rrule as {thm, elhs, ...}) =
526 ss |> map_simpset1 (fn (rules, prems, bounds, depth) =>
527 (Net.delete_term ((term_of elhs, rrule), rules, eq_rrule), prems, bounds, depth))
528 handle Net.DELETE => (warn_thm "Rewrite rule not in simpset:" thm; ss);
530 fun ss delsimps thms =
531 orient_comb_simps del_rrule (map mk_rrule2 o mk_rrule ss) (ss, thms);
536 fun cong_name (Const (a, _)) = Some a
537 | cong_name (Free (a, _)) = Some ("Free: " ^ a)
538 | cong_name _ = None;
542 fun is_full_cong_prems [] [] = true
543 | is_full_cong_prems [] _ = false
544 | is_full_cong_prems (p :: prems) varpairs =
545 (case Logic.strip_assums_concl p of
546 Const ("==", _) $ lhs $ rhs =>
547 let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in
548 is_Var x andalso forall is_Bound xs andalso
549 null (findrep xs) andalso xs = ys andalso
550 (x, y) mem varpairs andalso
551 is_full_cong_prems prems (varpairs \ (x, y))
555 fun is_full_cong thm =
557 val prems = prems_of thm and concl = concl_of thm;
558 val (lhs, rhs) = Logic.dest_equals concl;
559 val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs;
561 f = g andalso null (findrep (xs @ ys)) andalso length xs = length ys andalso
562 is_full_cong_prems prems (xs ~~ ys)
565 fun add_cong (ss, thm) = ss |>
566 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
568 val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm))
569 handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm);
570 (*val lhs = Pattern.eta_contract lhs;*)
571 val a = the (cong_name (head_of (term_of lhs))) handle Library.OPTION =>
572 raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm);
573 val (alist, weak) = congs;
574 val alist2 = overwrite_warn (alist, (a, {lhs = lhs, thm = thm}))
575 ("Overwriting congruence rule for " ^ quote a);
576 val weak2 = if is_full_cong thm then weak else a :: weak;
577 in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);
579 fun del_cong (ss, thm) = ss |>
580 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
582 val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ =>
583 raise SIMPLIFIER ("Congruence not a meta-equality", thm);
584 (*val lhs = Pattern.eta_contract lhs;*)
585 val a = the (cong_name (head_of lhs)) handle Library.OPTION =>
586 raise SIMPLIFIER ("Congruence must start with a constant", thm);
587 val (alist, _) = congs;
588 val alist2 = filter (fn (x, _) => x <> a) alist;
589 val weak2 = alist2 |> mapfilter (fn (a, {thm, ...}) =>
590 if is_full_cong thm then None else Some a);
591 in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);
593 fun mk_cong (Simpset (_, {mk_rews = {mk_cong = f, ...}, ...})) = f;
597 val (op addeqcongs) = foldl add_cong;
598 val (op deleqcongs) = foldl del_cong;
600 fun ss addcongs congs = ss addeqcongs map (mk_cong ss) congs;
601 fun ss delcongs congs = ss deleqcongs map (mk_cong ss) congs;
610 fun add_proc (ss, proc as Proc {name, lhs, ...}) =
611 (trace_cterm false ("Adding simplification procedure " ^ quote name ^ " for") lhs;
612 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
613 (congs, Net.insert_term ((term_of lhs, proc), procs, eq_proc),
614 mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss
616 (warning ("Ignoring duplicate simplification procedure " ^ quote name); ss));
618 fun del_proc (ss, proc as Proc {name, lhs, ...}) =
619 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
620 (congs, Net.delete_term ((term_of lhs, proc), procs, eq_proc),
621 mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss
623 (warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss);
627 val (op addsimprocs) = foldl (fn (ss, Simproc procs) => foldl add_proc (ss, procs));
628 val (op delsimprocs) = foldl (fn (ss, Simproc procs) => foldl del_proc (ss, procs));
637 fun map_mk_rews f = map_simpset2 (fn (congs, procs, {mk, mk2, mk_cong, mk_sym, mk_eq_True},
638 termless, subgoal_tac, loop_tacs, solvers) =>
639 let val (mk', mk2', mk_cong', mk_sym', mk_eq_True') = f (mk, mk2, mk_cong, mk_sym, mk_eq_True) in
640 (congs, procs, {mk = mk', mk2 = mk2', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True'},
641 termless, subgoal_tac, loop_tacs, solvers)
646 fun ss setmksimps mk = ss |> map_mk_rews (fn (_, mk2, mk_cong, mk_sym, mk_eq_True) =>
647 (mk, mk2, mk_cong, mk_sym, mk_eq_True));
649 fun ss setmksimps2 mk2 = ss |> map_mk_rews (fn (mk, _, mk_cong, mk_sym, mk_eq_True) =>
650 (mk, mk2, mk_cong, mk_sym, mk_eq_True));
652 fun ss setmkcong mk_cong = ss |> map_mk_rews (fn (mk, mk2, _, mk_sym, mk_eq_True) =>
653 (mk, mk2, mk_cong, mk_sym, mk_eq_True));
655 fun ss setmksym mk_sym = ss |> map_mk_rews (fn (mk, mk2, mk_cong, _, mk_eq_True) =>
656 (mk, mk2, mk_cong, mk_sym, mk_eq_True));
658 fun ss setmkeqTrue mk_eq_True = ss |> map_mk_rews (fn (mk, mk2, mk_cong, mk_sym, _) =>
659 (mk, mk2, mk_cong, mk_sym, mk_eq_True));
666 fun ss settermless termless = ss |>
667 map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) =>
668 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
673 fun ss setsubgoaler subgoal_tac = ss |>
674 map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) =>
675 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
677 fun ss setloop tac = ss |>
678 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) =>
679 (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers));
681 fun ss addloop (name, tac) = ss |>
682 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
683 (congs, procs, mk_rews, termless, subgoal_tac,
684 overwrite_warn (loop_tacs, (name, tac)) ("Overwriting looper " ^ quote name),
687 fun ss delloop name = ss |>
688 map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
689 let val loop_tacs' = filter_out (equal name o #1) loop_tacs in
690 if length loop_tacs <> length loop_tacs' then ()
691 else warning ("No such looper in simpset: " ^ quote name);
692 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs', solvers)
695 fun ss setSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless,
696 subgoal_tac, loop_tacs, (unsafe_solvers, _)) =>
697 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver])));
699 fun ss addSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless,
700 subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,
701 subgoal_tac, loop_tacs, (unsafe_solvers, merge_solvers solvers [solver])));
703 fun ss setSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless,
704 subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless,
705 subgoal_tac, loop_tacs, ([solver], solvers)));
707 fun ss addSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless,
708 subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,
709 subgoal_tac, loop_tacs, (merge_solvers unsafe_solvers [solver], solvers)));
711 fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, termless,
712 subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, termless,
713 subgoal_tac, loop_tacs, (solvers, solvers)));
720 Uses conversions, see:
721 L C Paulson, A higher-order implementation of rewriting,
722 Science of Computer Programming 3 (1983), pages 119-149.
725 val dest_eq = Drule.dest_equals o Thm.cprop_of;
726 val lhs_of = #1 o dest_eq;
727 val rhs_of = #2 o dest_eq;
729 fun check_conv msg thm thm' =
731 val thm'' = transitive thm (transitive
732 (symmetric (Drule.beta_eta_conversion (lhs_of thm'))) thm')
733 in if msg then trace_thm "SUCCEEDED" thm' else (); Some thm'' end
735 let val {sign, prop = _ $ _ $ prop0, ...} = Thm.rep_thm thm in
736 trace_thm "Proved wrong thm (Check subgoaler?)" thm';
737 trace_term false "Should have proved:" sign prop0;
744 fun mk_procrule thm =
745 let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in
746 if rewrite_rule_extra_vars prems lhs rhs
747 then (warn_thm "Extra vars on rhs:" thm; [])
748 else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}]
752 (* rewritec: conversion to apply the meta simpset to a term *)
754 (*Since the rewriting strategy is bottom-up, we avoid re-normalizing already
755 normalized terms by carrying around the rhs of the rewrite rule just
756 applied. This is called the `skeleton'. It is decomposed in parallel
757 with the term. Once a Var is encountered, the corresponding term is
758 already in normal form.
759 skel0 is a dummy skeleton that is to enforce complete normalization.*)
763 (*Use rhs as skeleton only if the lhs does not contain unnormalized bits.
764 The latter may happen iff there are weak congruence rules for constants
767 fun uncond_skel ((_, weak), (lhs, rhs)) =
768 if null weak then rhs (*optimization*)
769 else if exists_Const (fn (c, _) => c mem weak) lhs then skel0
772 (*Behaves like unconditional rule if rhs does not contain vars not in the lhs.
773 Otherwise those vars may become instantiated with unnormalized terms
774 while the premises are solved.*)
776 fun cond_skel (args as (congs, (lhs, rhs))) =
777 if term_varnames rhs subset term_varnames lhs then uncond_skel args
782 val ss' = ss |> map_simpset1 (fn (rules, prems, bounds, depth) =>
783 (rules, prems, bounds, depth + 1));
784 val Simpset ({depth = depth', ...}, _) = ss';
786 if depth' > ! simp_depth_limit
787 then (warning "simp_depth_limit exceeded - giving up"; None)
789 (if depth' mod 10 = 0
790 then warning ("Simplification depth " ^ string_of_int depth')
796 Rewriting -- we try in order:
798 (2) unconditional rewrite rules
799 (3) conditional rewrite rules
800 (4) simplification procedures
802 IMPORTANT: rewrite rules must not introduce new Vars or TVars!
805 fun rewritec (prover, signt, maxt) ss t =
807 val Simpset ({rules, ...}, {congs, procs, termless, ...}) = ss;
808 val eta_thm = Thm.eta_conversion t;
809 val eta_t' = rhs_of eta_thm;
810 val eta_t = term_of eta_t';
811 val tsigt = Sign.tsig_of signt;
812 fun rew {thm, name, lhs, elhs, fo, perm} =
814 val {sign, prop, maxidx, ...} = rep_thm thm;
815 val _ = if Sign.subsig (sign, signt) then ()
816 else (warn_thm "Ignoring rewrite rule from different theory:" thm;
817 raise Pattern.MATCH);
818 val (rthm, elhs') = if maxt = ~1 then (thm, elhs)
819 else (Thm.incr_indexes (maxt+1) thm, Thm.cterm_incr_indexes (maxt+1) elhs);
820 val insts = if fo then Thm.cterm_first_order_match (elhs', eta_t')
821 else Thm.cterm_match (elhs', eta_t');
822 val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm);
823 val prop' = Thm.prop_of thm';
824 val unconditional = (Logic.count_prems (prop',0) = 0);
825 val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop')
827 if perm andalso not (termless (rhs', lhs'))
828 then (trace_named_thm "Cannot apply permutative rewrite rule" (thm, name);
829 trace_thm "Term does not become smaller:" thm'; None)
830 else (trace_named_thm "Applying instance of rewrite rule" (thm, name);
833 (trace_thm "Rewriting:" thm';
834 let val lr = Logic.dest_equals prop;
835 val Some thm'' = check_conv false eta_thm thm'
836 in Some (thm'', uncond_skel (congs, lr)) end)
838 (trace_thm "Trying to rewrite:" thm';
839 case incr_depth ss of
840 None => (trace_thm "FAILED - reached depth limit" thm'; None)
842 (case prover ss' thm' of
843 None => (trace_thm "FAILED" thm'; None)
845 (case check_conv true eta_thm thm2 of
848 let val concl = Logic.strip_imp_concl prop
849 val lr = Logic.dest_equals concl
850 in Some (thm2', cond_skel (congs, lr)) end))))
854 | rews (rrule :: rrules) =
855 let val opt = rew rrule handle Pattern.MATCH => None
856 in case opt of None => rews rrules | some => some end;
858 fun sort_rrules rrs = let
859 fun is_simple({thm, ...}:rrule) = case Thm.prop_of thm of
860 Const("==",_) $ _ $ _ => true
862 fun sort [] (re1,re2) = re1 @ re2
863 | sort (rr::rrs) (re1,re2) = if is_simple rr
864 then sort rrs (rr::re1,re2)
865 else sort rrs (re1,rr::re2)
866 in sort rrs ([],[]) end
868 fun proc_rews [] = None
869 | proc_rews (Proc {name, proc, lhs, ...} :: ps) =
870 if Pattern.matches tsigt (Thm.term_of lhs, Thm.term_of t) then
871 (debug_term false ("Trying procedure " ^ quote name ^ " on:") signt eta_t;
872 case transform_failure (curry SIMPROC_FAIL name)
873 (fn () => proc signt ss eta_t) () of
874 None => (debug false "FAILED"; proc_rews ps)
876 (trace_thm ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm;
877 (case rews (mk_procrule raw_thm) of
878 None => (trace_cterm true ("IGNORED result of simproc " ^ quote name ^
879 " -- does not match") t; proc_rews ps)
883 Abs _ $ _ => Some (transitive eta_thm
884 (beta_conversion false eta_t'), skel0)
885 | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of
886 None => proc_rews (Net.match_term procs eta_t)
891 (* conversion to apply a congruence rule to a term *)
893 fun congc (prover,signt,maxt) {thm=cong,lhs=lhs} t =
894 let val sign = Thm.sign_of_thm cong
895 val _ = if Sign.subsig (sign, signt) then ()
896 else error("Congruence rule from different theory")
897 val rthm = if maxt = ~1 then cong else Thm.incr_indexes (maxt+1) cong;
898 val rlhs = fst (Drule.dest_equals (Drule.strip_imp_concl (cprop_of rthm)));
899 val insts = Thm.cterm_match (rlhs, t)
900 (* Pattern.match can raise Pattern.MATCH;
901 is handled when congc is called *)
902 val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
903 val unit = trace_thm "Applying congruence rule:" thm';
904 fun err (msg, thm) = (trace_thm msg thm; None)
905 in case prover thm' of
906 None => err ("Congruence proof failed. Could not prove", thm')
907 | Some thm2 => (case check_conv true (Drule.beta_eta_conversion t) thm2 of
908 None => err ("Congruence proof failed. Should not have proved", thm2)
910 if op aconv (pairself term_of (dest_equals (cprop_of thm2')))
911 then None else Some thm2')
915 apsnd dest_equals (dest_implies (hd (cprems_of Drule.imp_cong)));
917 fun transitive1 None None = None
918 | transitive1 (Some thm1) None = Some thm1
919 | transitive1 None (Some thm2) = Some thm2
920 | transitive1 (Some thm1) (Some thm2) = Some (transitive thm1 thm2)
922 fun transitive2 thm = transitive1 (Some thm);
923 fun transitive3 thm = transitive1 thm o Some;
925 fun bottomc ((simprem, useprem, mutsimp), prover, sign, maxidx) =
928 if is_Var skel then None
930 (case subc skel ss t of
932 (case rewritec (prover, sign, maxidx) ss (rhs_of thm1) of
933 Some (thm2, skel2) =>
934 transitive2 (transitive thm1 thm2)
935 (botc skel2 ss (rhs_of thm2))
938 (case rewritec (prover, sign, maxidx) ss t of
939 Some (thm2, skel2) => transitive2 thm2
940 (botc skel2 ss (rhs_of thm2))
944 (case botc skel0 ss t of
945 Some trec1 => trec1 | None => (reflexive t))
947 and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 =
951 val b = variant bounds a;
952 val (v, t') = Thm.dest_abs (Some ("." ^ b)) t0;
953 val ss' = add_bound b ss;
954 val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0;
955 in case botc skel' ss' t' of
956 Some thm => Some (abstract_rule a v thm)
959 | t $ _ => (case t of
960 Const ("==>", _) $ _ => impc t0 ss
962 let val thm = beta_conversion false t0
963 in case subc skel0 ss (rhs_of thm) of
965 | Some thm' => Some (transitive thm thm')
970 val (tskel, uskel) = case skel of
971 tskel $ uskel => (tskel, uskel)
972 | _ => (skel0, skel0);
973 val (ct, cu) = Thm.dest_comb t0
975 (case botc tskel ss ct of
977 (case botc uskel ss cu of
978 Some thm2 => Some (combination thm1 thm2)
979 | None => Some (combination thm1 (reflexive cu)))
981 (case botc uskel ss cu of
982 Some thm1 => Some (combination (reflexive ct) thm1)
985 val (h, ts) = strip_comb t
986 in case cong_name h of
988 (case assoc_string (fst congs, a) of
991 (*post processing: some partial applications h t1 ... tj, j <= length ts,
992 may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*)
994 val thm = congc (prover ss, sign, maxidx) cong t0;
995 val t = if_none (apsome rhs_of thm) t0;
996 val (cl, cr) = Thm.dest_comb t
997 val dVar = Var(("", 0), dummyT)
999 list_comb (h, replicate (length ts) dVar)
1000 in case botc skel ss cl of
1002 | Some thm' => transitive3 thm
1003 (combination thm' (reflexive cr))
1004 end handle TERM _ => error "congc result"
1005 | Pattern.MATCH => appc ()))
1011 if mutsimp then mut_impc0 [] ct [] [] ss else nonmut_impc ct ss
1013 and rules_of_prem ss prem =
1014 if maxidx_of_term (term_of prem) <> ~1
1015 then (trace_cterm true
1016 "Cannot add premise as rewrite rule because it contains (type) unknowns:" prem; ([], None))
1018 let val asm = assume prem
1019 in (extract_safe_rrules (ss, asm), Some asm) end
1021 and add_rrules (rrss, asms) ss =
1022 let val Asms = mapfilter I asms
1023 in foldl (insert_rrule true) (ss, flat rrss) |> add_prems2 Asms end
1025 and disch r (prem, eq) =
1027 val (lhs, rhs) = dest_eq eq;
1028 val eq' = implies_elim (Thm.instantiate
1029 ([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong)
1030 (implies_intr prem eq)
1031 in if not r then eq' else
1033 val (prem', concl) = dest_implies lhs;
1034 val (prem'', _) = dest_implies rhs
1035 in transitive (transitive
1036 (Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)])
1037 Drule.swap_prems_eq) eq')
1038 (Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)])
1039 Drule.swap_prems_eq)
1043 and rebuild [] _ _ _ _ eq = eq
1044 | rebuild (prem :: prems) concl (rrs :: rrss) (asm :: asms) ss eq =
1046 val ss' = add_rrules (rev rrss, rev asms) ss;
1048 Drule.mk_implies (prem, if_none (apsome rhs_of eq) concl);
1049 val dprem = apsome (curry (disch false) prem)
1050 in case rewritec (prover, sign, maxidx) ss' concl' of
1051 None => rebuild prems concl' rrss asms ss (dprem eq)
1052 | Some (eq', _) => transitive2 (foldl (disch false o swap)
1053 (the (transitive3 (dprem eq) eq'), prems))
1054 (mut_impc0 (rev prems) (rhs_of eq') (rev rrss) (rev asms) ss)
1057 and mut_impc0 prems concl rrss asms ss =
1059 val prems' = strip_imp_prems concl;
1060 val (rrss', asms') = split_list (map (rules_of_prem ss) prems')
1061 in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')
1062 (asms @ asms') [] [] [] [] ss ~1 ~1
1065 and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k =
1066 transitive1 (foldl (fn (eq2, (eq1, prem)) => transitive1 eq1
1067 (apsome (curry (disch false) prem) eq2)) (None, eqns ~~ prems'))
1068 (if changed > 0 then
1069 mut_impc (rev prems') concl (rev rrss') (rev asms')
1070 [] [] [] [] ss ~1 changed
1071 else rebuild prems' concl rrss' asms' ss
1072 (botc skel0 (add_rrules (rev rrss', rev asms') ss) concl))
1074 | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms)
1075 prems' rrss' asms' eqns ss changed k =
1076 case (if k = 0 then None else botc skel0 (add_rrules
1077 (rev rrss' @ rrss, rev asms' @ asms) ss) prem) of
1078 None => mut_impc prems concl rrss asms (prem :: prems')
1079 (rrs :: rrss') (asm :: asms') (None :: eqns) ss changed
1080 (if k = 0 then 0 else k - 1)
1083 val prem' = rhs_of eqn;
1084 val tprems = map term_of prems;
1085 val i = 1 + foldl Int.max (~1, map (fn p =>
1086 find_index_eq p tprems) (#hyps (rep_thm eqn)));
1087 val (rrs', asm') = rules_of_prem ss prem'
1088 in mut_impc prems concl rrss asms (prem' :: prems')
1089 (rrs' :: rrss') (asm' :: asms') (Some (foldr (disch true)
1090 (take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies
1091 (drop (i, prems), concl))))) :: eqns) ss (length prems') ~1
1094 (*legacy code - only for backwards compatibility*)
1095 and nonmut_impc ct ss =
1096 let val (prem, conc) = dest_implies ct;
1097 val thm1 = if simprem then botc skel0 ss prem else None;
1098 val prem1 = if_none (apsome rhs_of thm1) prem;
1099 val ss1 = if not useprem then ss else add_rrules
1100 (apsnd single (apfst single (rules_of_prem ss prem1))) ss
1101 in (case botc skel0 ss1 conc of
1102 None => (case thm1 of
1104 | Some thm1' => Some (Drule.imp_cong' thm1' (reflexive conc)))
1106 let val thm2' = disch false (prem1, thm2)
1110 Some (transitive (Drule.imp_cong' thm1' (reflexive conc)) thm2'))
1117 (* Meta-rewriting: rewrites t to u and returns the theorem t==u *)
1122 use A in simplifying B,
1123 use prems of B (if B is again a meta-impl.) to simplify A)
1124 when simplifying A ==> B
1125 prover: how to solve premises in conditional rewrites and congruences
1128 fun rewrite_cterm mode prover ss ct =
1130 val Simpset ({depth, ...}, _) = ss;
1131 val {sign, t, maxidx, ...} = Thm.rep_cterm ct;
1133 trace_cterm false "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" ct;
1134 simp_depth := depth;
1135 bottomc (mode, prover, sign, maxidx) ss ct
1136 end handle THM (s, _, thms) =>
1137 error ("Exception THM was raised in simplifier:\n" ^ s ^ "\n" ^
1138 Pretty.string_of (Display.pretty_thms thms));
1141 fun rewrite_aux _ _ [] = (fn ct => Thm.reflexive ct)
1142 | rewrite_aux prover full thms =
1143 rewrite_cterm (full, false, false) prover (empty_ss addsimps thms);
1145 (*Rewrite a theorem*)
1146 fun simplify_aux _ _ [] = (fn th => th)
1147 | simplify_aux prover full thms =
1148 Drule.fconv_rule (rewrite_cterm (full, false, false) prover (empty_ss addsimps thms));
1150 (*simple term rewriting -- no proof*)
1151 fun rewrite_term sg rules procs =
1152 Pattern.rewrite_term (Sign.tsig_of sg) (map decomp_simp' rules) procs;
1154 fun rewrite_thm mode prover ss = Drule.fconv_rule (rewrite_cterm mode prover ss);
1156 (*Rewrite the subgoals of a proof state (represented by a theorem) *)
1157 fun rewrite_goals_rule_aux _ [] th = th
1158 | rewrite_goals_rule_aux prover thms th =
1159 Drule.fconv_rule (Drule.goals_conv (K true) (rewrite_cterm (true, true, false) prover
1160 (empty_ss addsimps thms))) th;
1162 (*Rewrite the subgoal of a proof state (represented by a theorem)*)
1163 fun rewrite_goal_rule mode prover ss i thm =
1164 if 0 < i andalso i <= nprems_of thm
1165 then Drule.fconv_rule (Drule.goals_conv (fn j => j=i) (rewrite_cterm mode prover ss)) thm
1166 else raise THM("rewrite_goal_rule",i,[thm]);
1168 (*Rewrite subgoal i only. SELECT_GOAL avoids inefficiencies in goals_conv.*)
1169 fun asm_rewrite_goal_tac mode prover_tac ss =
1171 (PRIMITIVE (rewrite_goal_rule mode (SINGLE o prover_tac) ss 1));
1175 (** simplification tactics and rules **)
1177 fun solve_all_tac solvers ss =
1179 val Simpset (_, {subgoal_tac, ...}) = ss;
1180 val solve_tac = subgoal_tac (set_solvers solvers ss) THEN_ALL_NEW (K no_tac);
1181 in DEPTH_SOLVE (solve_tac 1) end;
1183 (*NOTE: may instantiate unknowns that appear also in other subgoals*)
1184 fun generic_simp_tac safe mode ss =
1186 val Simpset ({prems, ...}, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = ss;
1187 val loop_tac = FIRST' (map #2 loop_tacs);
1188 val solve_tac = FIRST' (map (solver prems) (if safe then solvers else unsafe_solvers));
1190 fun simp_loop_tac i =
1191 asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN
1192 (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));
1193 in simp_loop_tac end;
1195 fun simp rew mode ss thm =
1197 val Simpset (_, {solvers = (unsafe_solvers, _), ...}) = ss;
1198 val tacf = solve_all_tac unsafe_solvers;
1199 fun prover s th = apsome #1 (Seq.pull (tacf s th));
1200 in rew mode prover ss thm end;
1202 val simp_thm = simp rewrite_thm;
1203 val simp_cterm = simp rewrite_cterm;
1207 structure BasicMetaSimplifier: BASIC_META_SIMPLIFIER = MetaSimplifier;
1208 open BasicMetaSimplifier;