berghofe@10413: (* Title: Pure/meta_simplifier.ML wenzelm@29269: Author: Tobias Nipkow and Stefan Berghofer, TU Muenchen berghofe@10413: wenzelm@11672: Meta-level Simplification. berghofe@10413: *) berghofe@10413: skalberg@15006: infix 4 wenzelm@15023: addsimps delsimps addeqcongs deleqcongs addcongs delcongs addsimprocs delsimprocs nipkow@15199: setmksimps setmkcong setmksym setmkeqTrue settermless setsubgoaler wenzelm@17882: setloop' setloop addloop addloop' delloop setSSolver addSSolver setSolver addSolver; skalberg@15006: wenzelm@11672: signature BASIC_META_SIMPLIFIER = wenzelm@11672: sig wenzelm@32738: val debug_simp: bool Unsynchronized.ref wenzelm@32738: val trace_simp: bool Unsynchronized.ref wenzelm@32738: val trace_simp_depth_limit: int Unsynchronized.ref wenzelm@15023: type rrule wenzelm@16807: val eq_rrule: rrule * rrule -> bool wenzelm@15023: type simpset wenzelm@15023: type proc wenzelm@17614: type solver wenzelm@17614: val mk_solver': string -> (simpset -> int -> tactic) -> solver wenzelm@17614: val mk_solver: string -> (thm list -> int -> tactic) -> solver wenzelm@15023: val empty_ss: simpset wenzelm@15023: val merge_ss: simpset * simpset -> simpset wenzelm@30356: val dest_ss: simpset -> wenzelm@30356: {simps: (string * thm) list, wenzelm@30356: procs: (string * cterm list) list, wenzelm@30356: congs: (string * thm) list, wenzelm@30356: weak_congs: string list, wenzelm@30356: loopers: string list, wenzelm@30356: unsafe_solvers: string list, wenzelm@30356: safe_solvers: string list} wenzelm@15023: type simproc wenzelm@22234: val eq_simproc: simproc * simproc -> bool wenzelm@22234: val morph_simproc: morphism -> simproc -> simproc wenzelm@22234: val make_simproc: {name: string, lhss: cterm list, wenzelm@22234: proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> simproc wenzelm@22008: val mk_simproc: string -> cterm list -> (theory -> simpset -> term -> thm option) -> simproc wenzelm@15023: val add_prems: thm list -> simpset -> simpset wenzelm@15023: val prems_of_ss: simpset -> thm list wenzelm@15023: val addsimps: simpset * thm list -> simpset wenzelm@15023: val delsimps: simpset * thm list -> simpset wenzelm@15023: val addeqcongs: simpset * thm list -> simpset wenzelm@15023: val deleqcongs: simpset * thm list -> simpset wenzelm@15023: val addcongs: simpset * thm list -> simpset wenzelm@15023: val delcongs: simpset * thm list -> simpset wenzelm@15023: val addsimprocs: simpset * simproc list -> simpset wenzelm@15023: val delsimprocs: simpset * simproc list -> simpset wenzelm@30318: val mksimps: simpset -> thm -> thm list wenzelm@15023: val setmksimps: simpset * (thm -> thm list) -> simpset wenzelm@15023: val setmkcong: simpset * (thm -> thm) -> simpset wenzelm@15023: val setmksym: simpset * (thm -> thm option) -> simpset wenzelm@15023: val setmkeqTrue: simpset * (thm -> thm option) -> simpset wenzelm@15023: val settermless: simpset * (term * term -> bool) -> simpset wenzelm@15023: val setsubgoaler: simpset * (simpset -> int -> tactic) -> simpset wenzelm@17882: val setloop': simpset * (simpset -> int -> tactic) -> simpset wenzelm@15023: val setloop: simpset * (int -> tactic) -> simpset wenzelm@17882: val addloop': simpset * (string * (simpset -> int -> tactic)) -> simpset wenzelm@15023: val addloop: simpset * (string * (int -> tactic)) -> simpset wenzelm@15023: val delloop: simpset * string -> simpset wenzelm@15023: val setSSolver: simpset * solver -> simpset wenzelm@15023: val addSSolver: simpset * solver -> simpset wenzelm@15023: val setSolver: simpset * solver -> simpset wenzelm@15023: val addSolver: simpset * solver -> simpset wenzelm@21708: wenzelm@21708: val rewrite_rule: thm list -> thm -> thm wenzelm@21708: val rewrite_goals_rule: thm list -> thm -> thm wenzelm@21708: val rewrite_goals_tac: thm list -> tactic wenzelm@23536: val rewrite_goal_tac: thm list -> int -> tactic wenzelm@21708: val rewtac: thm -> tactic wenzelm@21708: val prune_params_tac: tactic wenzelm@21708: val fold_rule: thm list -> thm -> thm wenzelm@21708: val fold_goals_tac: thm list -> tactic wenzelm@30552: val norm_hhf: thm -> thm wenzelm@30552: val norm_hhf_protect: thm -> thm skalberg@15006: end; skalberg@15006: berghofe@10413: signature META_SIMPLIFIER = berghofe@10413: sig wenzelm@11672: include BASIC_META_SIMPLIFIER berghofe@10413: exception SIMPLIFIER of string * thm wenzelm@30336: val internal_ss: simpset -> wenzelm@30336: {rules: rrule Net.net, wenzelm@30336: prems: thm list, wenzelm@30336: bounds: int * ((string * typ) * string) list, wenzelm@32738: depth: int * bool Unsynchronized.ref, wenzelm@30336: context: Proof.context option} * krauss@30908: {congs: (string * thm) list * string list, wenzelm@30336: procs: proc Net.net, wenzelm@30336: mk_rews: wenzelm@30336: {mk: thm -> thm list, wenzelm@30336: mk_cong: thm -> thm, wenzelm@30336: mk_sym: thm -> thm option, wenzelm@30336: mk_eq_True: thm -> thm option, wenzelm@30336: reorient: theory -> term list -> term -> term -> bool}, wenzelm@30336: termless: term * term -> bool, wenzelm@30336: subgoal_tac: simpset -> int -> tactic, wenzelm@30336: loop_tacs: (string * (simpset -> int -> tactic)) list, wenzelm@30336: solvers: solver list * solver list} haftmann@27558: val add_simp: thm -> simpset -> simpset haftmann@27558: val del_simp: thm -> simpset -> simpset wenzelm@17966: val solver: simpset -> solver -> int -> tactic wenzelm@24124: val simp_depth_limit_value: Config.value Config.T wenzelm@24124: val simp_depth_limit: int Config.T wenzelm@15023: val clear_ss: simpset -> simpset wenzelm@16458: val simproc_i: theory -> string -> term list wenzelm@16458: -> (theory -> simpset -> term -> thm option) -> simproc wenzelm@16458: val simproc: theory -> string -> string list wenzelm@16458: -> (theory -> simpset -> term -> thm option) -> simproc wenzelm@17882: val inherit_context: simpset -> simpset -> simpset wenzelm@20289: val the_context: simpset -> Proof.context wenzelm@20289: val context: Proof.context -> simpset -> simpset wenzelm@17897: val theory_context: theory -> simpset -> simpset wenzelm@32738: val debug_bounds: bool Unsynchronized.ref wenzelm@18208: val set_reorient: (theory -> term list -> term -> term -> bool) -> simpset -> simpset wenzelm@17966: val set_solvers: solver list -> simpset -> simpset wenzelm@23598: val rewrite_cterm: bool * bool * bool -> (simpset -> thm -> thm option) -> simpset -> conv wenzelm@16458: val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term wenzelm@15023: val rewrite_thm: bool * bool * bool -> wenzelm@15023: (simpset -> thm -> thm option) -> simpset -> thm -> thm wenzelm@15023: val rewrite_goal_rule: bool * bool * bool -> wenzelm@15023: (simpset -> thm -> thm option) -> simpset -> int -> thm -> thm wenzelm@23536: val asm_rewrite_goal_tac: bool * bool * bool -> wenzelm@23536: (simpset -> tactic) -> simpset -> int -> tactic wenzelm@23598: val rewrite: bool -> thm list -> conv wenzelm@21708: val simplify: bool -> thm list -> thm -> thm berghofe@10413: end; berghofe@10413: wenzelm@15023: structure MetaSimplifier: META_SIMPLIFIER = berghofe@10413: struct berghofe@10413: wenzelm@15023: (** datatype simpset **) wenzelm@15023: wenzelm@15023: (* rewrite rules *) berghofe@10413: wenzelm@20546: type rrule = wenzelm@20546: {thm: thm, (*the rewrite rule*) wenzelm@20546: name: string, (*name of theorem from which rewrite rule was extracted*) wenzelm@20546: lhs: term, (*the left-hand side*) wenzelm@20546: elhs: cterm, (*the etac-contracted lhs*) wenzelm@20546: extra: bool, (*extra variables outside of elhs*) wenzelm@20546: fo: bool, (*use first-order matching*) wenzelm@20546: perm: bool}; (*the rewrite rule is permutative*) wenzelm@15023: wenzelm@20546: (* wenzelm@12603: Remarks: berghofe@10413: - elhs is used for matching, wenzelm@15023: lhs only for preservation of bound variable names; berghofe@10413: - fo is set iff berghofe@10413: either elhs is first-order (no Var is applied), wenzelm@15023: in which case fo-matching is complete, berghofe@10413: or elhs is not a pattern, wenzelm@20546: in which case there is nothing better to do; wenzelm@20546: *) berghofe@10413: berghofe@10413: fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = wenzelm@22360: Thm.eq_thm_prop (thm1, thm2); wenzelm@15023: wenzelm@15023: wenzelm@17614: (* simplification sets, procedures, and solvers *) wenzelm@15023: wenzelm@15023: (*A simpset contains data required during conversion: berghofe@10413: rules: discrimination net of rewrite rules; wenzelm@15023: prems: current premises; berghofe@15249: bounds: maximal index of bound variables already used wenzelm@15023: (for generating new names when rewriting under lambda abstractions); wenzelm@22892: depth: simp_depth and exceeded flag; berghofe@10413: congs: association list of congruence rules and berghofe@10413: a list of `weak' congruence constants. berghofe@10413: A congruence is `weak' if it avoids normalization of some argument. berghofe@10413: procs: discrimination net of simplification procedures berghofe@10413: (functions that prove rewrite rules on the fly); wenzelm@15023: mk_rews: wenzelm@15023: mk: turn simplification thms into rewrite rules; wenzelm@15023: mk_cong: prepare congruence rules; wenzelm@15023: mk_sym: turn == around; wenzelm@15023: mk_eq_True: turn P into P == True; wenzelm@15023: termless: relation for ordered rewriting;*) skalberg@15011: wenzelm@15023: type mk_rews = wenzelm@15023: {mk: thm -> thm list, wenzelm@15023: mk_cong: thm -> thm, wenzelm@15023: mk_sym: thm -> thm option, wenzelm@18208: mk_eq_True: thm -> thm option, wenzelm@18208: reorient: theory -> term list -> term -> term -> bool}; wenzelm@15023: wenzelm@15023: datatype simpset = wenzelm@15023: Simpset of wenzelm@15023: {rules: rrule Net.net, berghofe@10413: prems: thm list, wenzelm@17882: bounds: int * ((string * typ) * string) list, wenzelm@32738: depth: int * bool Unsynchronized.ref, wenzelm@20289: context: Proof.context option} * krauss@30908: {congs: (string * thm) list * string list, wenzelm@15023: procs: proc Net.net, wenzelm@15023: mk_rews: mk_rews, nipkow@11504: termless: term * term -> bool, skalberg@15011: subgoal_tac: simpset -> int -> tactic, wenzelm@17882: loop_tacs: (string * (simpset -> int -> tactic)) list, wenzelm@15023: solvers: solver list * solver list} wenzelm@15023: and proc = wenzelm@15023: Proc of wenzelm@15023: {name: string, wenzelm@15023: lhs: cterm, wenzelm@22008: proc: simpset -> cterm -> thm option, wenzelm@22234: id: stamp * thm list} wenzelm@17614: and solver = wenzelm@17614: Solver of wenzelm@17614: {name: string, wenzelm@17614: solver: simpset -> int -> tactic, wenzelm@15023: id: stamp}; wenzelm@15023: wenzelm@15023: wenzelm@30336: fun internal_ss (Simpset args) = args; berghofe@10413: wenzelm@22892: fun make_ss1 (rules, prems, bounds, depth, context) = wenzelm@22892: {rules = rules, prems = prems, bounds = bounds, depth = depth, context = context}; wenzelm@15023: wenzelm@22892: fun map_ss1 f {rules, prems, bounds, depth, context} = wenzelm@22892: make_ss1 (f (rules, prems, bounds, depth, context)); berghofe@10413: wenzelm@15023: fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) = wenzelm@15023: {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless, wenzelm@15023: subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; wenzelm@15023: wenzelm@15023: fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} = wenzelm@15023: make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); wenzelm@15023: wenzelm@15023: fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); berghofe@10413: wenzelm@15023: fun map_simpset1 f (Simpset (r1, r2)) = Simpset (map_ss1 f r1, r2); wenzelm@15023: fun map_simpset2 f (Simpset (r1, r2)) = Simpset (r1, map_ss2 f r2); wenzelm@15023: wenzelm@17614: fun prems_of_ss (Simpset ({prems, ...}, _)) = prems; wenzelm@17614: wenzelm@22234: fun eq_procid ((s1: stamp, ths1: thm list), (s2, ths2)) = wenzelm@22360: s1 = s2 andalso eq_list Thm.eq_thm (ths1, ths2); wenzelm@22234: fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = eq_procid (id1, id2); wenzelm@17614: wenzelm@17614: fun mk_solver' name solver = Solver {name = name, solver = solver, id = stamp ()}; wenzelm@17614: fun mk_solver name solver = mk_solver' name (solver o prems_of_ss); wenzelm@17614: wenzelm@17614: fun solver_name (Solver {name, ...}) = name; wenzelm@17966: fun solver ss (Solver {solver = tac, ...}) = tac ss; wenzelm@17614: fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); wenzelm@17614: wenzelm@15023: wenzelm@22892: (* simp depth *) wenzelm@22892: wenzelm@24124: val simp_depth_limit_value = Config.declare false "simp_depth_limit" (Config.Int 100); wenzelm@24124: val simp_depth_limit = Config.int simp_depth_limit_value; wenzelm@24124: wenzelm@32738: val trace_simp_depth_limit = Unsynchronized.ref 1; wenzelm@22892: wenzelm@22892: fun trace_depth (Simpset ({depth = (depth, exceeded), ...}, _)) msg = wenzelm@23938: if depth > ! trace_simp_depth_limit then wenzelm@23938: if ! exceeded then () else (tracing "trace_simp_depth_limit exceeded!"; exceeded := true) wenzelm@22892: else wenzelm@23938: (tracing (enclose "[" "]" (string_of_int depth) ^ msg); exceeded := false); wenzelm@22892: wenzelm@22892: val inc_simp_depth = map_simpset1 (fn (rules, prems, bounds, (depth, exceeded), context) => wenzelm@22892: (rules, prems, bounds, wenzelm@32738: (depth + 1, wenzelm@32738: if depth = ! trace_simp_depth_limit then Unsynchronized.ref false else exceeded), context)); wenzelm@22892: wenzelm@22892: fun simp_depth (Simpset ({depth = (depth, _), ...}, _)) = depth; wenzelm@22892: wenzelm@22892: wenzelm@16985: (* diagnostics *) wenzelm@16985: wenzelm@16985: exception SIMPLIFIER of string * thm; wenzelm@16985: wenzelm@32738: val debug_simp = Unsynchronized.ref false; wenzelm@32738: val trace_simp = Unsynchronized.ref false; wenzelm@22892: wenzelm@16985: local wenzelm@16985: wenzelm@22892: fun prnt ss warn a = if warn then warning a else trace_depth ss a; wenzelm@16985: wenzelm@16985: fun show_bounds (Simpset ({bounds = (_, bs), ...}, _)) t = wenzelm@16985: let wenzelm@20146: val names = Term.declare_term_names t Name.context; wenzelm@20146: val xs = rev (#1 (Name.variants (rev (map #2 bs)) names)); wenzelm@17614: fun subst (((b, T), _), x') = (Free (b, T), Syntax.mark_boundT (x', T)); wenzelm@16985: in Term.subst_atomic (ListPair.map subst (bs, xs)) t end; wenzelm@16985: wenzelm@17705: in wenzelm@17705: wenzelm@22892: fun print_term ss warn a thy t = prnt ss warn (a ^ "\n" ^ wenzelm@26939: Syntax.string_of_term_global thy (if ! debug_simp then t else show_bounds ss t)); wenzelm@16985: wenzelm@22892: fun debug warn a ss = if ! debug_simp then prnt ss warn (a ()) else (); wenzelm@22892: fun trace warn a ss = if ! trace_simp then prnt ss warn (a ()) else (); wenzelm@16985: wenzelm@22892: fun debug_term warn a ss thy t = if ! debug_simp then print_term ss warn (a ()) thy t else (); wenzelm@22892: fun trace_term warn a ss thy t = if ! trace_simp then print_term ss warn (a ()) thy t else (); wenzelm@16985: wenzelm@16985: fun trace_cterm warn a ss ct = wenzelm@22892: if ! trace_simp then print_term ss warn (a ()) (Thm.theory_of_cterm ct) (Thm.term_of ct) wenzelm@22254: else (); wenzelm@16985: wenzelm@16985: fun trace_thm a ss th = wenzelm@22892: if ! trace_simp then print_term ss false (a ()) (Thm.theory_of_thm th) (Thm.full_prop_of th) wenzelm@22254: else (); wenzelm@16985: wenzelm@16985: fun trace_named_thm a ss (th, name) = wenzelm@16985: if ! trace_simp then wenzelm@22892: print_term ss false (if name = "" then a () else a () ^ " " ^ quote name ^ ":") wenzelm@16985: (Thm.theory_of_thm th) (Thm.full_prop_of th) wenzelm@16985: else (); wenzelm@16985: wenzelm@22892: fun warn_thm a ss th = wenzelm@22892: print_term ss true a (Thm.theory_of_thm th) (Thm.full_prop_of th); wenzelm@16985: wenzelm@20028: fun cond_warn_thm a (ss as Simpset ({context, ...}, _)) th = wenzelm@20546: if is_some context then () else warn_thm a ss th; wenzelm@20028: wenzelm@16985: end; wenzelm@16985: wenzelm@16985: berghofe@10413: berghofe@10413: (** simpset operations **) berghofe@10413: wenzelm@17882: (* context *) berghofe@10413: wenzelm@17614: fun eq_bound (x: string, (y, _)) = x = y; wenzelm@17614: wenzelm@22892: fun add_bound bound = map_simpset1 (fn (rules, prems, (count, bounds), depth, context) => wenzelm@22892: (rules, prems, (count + 1, bound :: bounds), depth, context)); wenzelm@17882: wenzelm@22892: fun add_prems ths = map_simpset1 (fn (rules, prems, bounds, depth, context) => wenzelm@22892: (rules, ths @ prems, bounds, depth, context)); wenzelm@17882: wenzelm@22892: fun inherit_context (Simpset ({bounds, depth, context, ...}, _)) = wenzelm@22892: map_simpset1 (fn (rules, prems, _, _, _) => (rules, prems, bounds, depth, context)); wenzelm@16985: wenzelm@17882: fun the_context (Simpset ({context = SOME ctxt, ...}, _)) = ctxt wenzelm@17882: | the_context _ = raise Fail "Simplifier: no proof context in simpset"; berghofe@10413: wenzelm@17897: fun context ctxt = wenzelm@22892: map_simpset1 (fn (rules, prems, bounds, depth, _) => (rules, prems, bounds, depth, SOME ctxt)); wenzelm@17882: wenzelm@21516: val theory_context = context o ProofContext.init; wenzelm@17897: wenzelm@27312: fun activate_context thy ss = wenzelm@27312: let wenzelm@27312: val ctxt = the_context ss; wenzelm@33031: val ctxt' = Context.raw_transfer (Theory.merge (thy, ProofContext.theory_of ctxt)) ctxt; wenzelm@27312: in context ctxt' ss end; wenzelm@17897: wenzelm@17897: wenzelm@20028: (* maintain simp rules *) berghofe@10413: wenzelm@20546: (* FIXME: it seems that the conditions on extra variables are too liberal if wenzelm@20546: prems are nonempty: does solving the prems really guarantee instantiation of wenzelm@20546: all its Vars? Better: a dynamic check each time a rule is applied. wenzelm@20546: *) wenzelm@20546: fun rewrite_rule_extra_vars prems elhs erhs = wenzelm@20546: let wenzelm@20546: val elhss = elhs :: prems; wenzelm@20546: val tvars = fold Term.add_tvars elhss []; wenzelm@20546: val vars = fold Term.add_vars elhss []; wenzelm@20546: in wenzelm@20546: erhs |> Term.exists_type (Term.exists_subtype wenzelm@20546: (fn TVar v => not (member (op =) tvars v) | _ => false)) orelse wenzelm@20546: erhs |> Term.exists_subterm wenzelm@20546: (fn Var v => not (member (op =) vars v) | _ => false) wenzelm@20546: end; wenzelm@20546: wenzelm@20546: fun rrule_extra_vars elhs thm = wenzelm@20546: rewrite_rule_extra_vars [] (term_of elhs) (Thm.full_prop_of thm); wenzelm@20546: wenzelm@15023: fun mk_rrule2 {thm, name, lhs, elhs, perm} = wenzelm@15023: let wenzelm@20546: val t = term_of elhs; wenzelm@20546: val fo = Pattern.first_order t orelse not (Pattern.pattern t); wenzelm@20546: val extra = rrule_extra_vars elhs thm; wenzelm@20546: in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end; berghofe@10413: wenzelm@20028: fun del_rrule (rrule as {thm, elhs, ...}) ss = wenzelm@22892: ss |> map_simpset1 (fn (rules, prems, bounds, depth, context) => wenzelm@22892: (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, bounds, depth, context)) wenzelm@20028: handle Net.DELETE => (cond_warn_thm "Rewrite rule not in simpset:" ss thm; ss); wenzelm@20028: wenzelm@32797: fun insert_rrule (rrule as {thm, name, ...}) ss = wenzelm@22254: (trace_named_thm (fn () => "Adding rewrite rule") ss (thm, name); wenzelm@22892: ss |> map_simpset1 (fn (rules, prems, bounds, depth, context) => wenzelm@15023: let wenzelm@15023: val rrule2 as {elhs, ...} = mk_rrule2 rrule; wenzelm@16807: val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules; wenzelm@22892: in (rules', prems, bounds, depth, context) end) wenzelm@20028: handle Net.INSERT => (cond_warn_thm "Ignoring duplicate rewrite rule:" ss thm; ss)); berghofe@10413: berghofe@10413: fun vperm (Var _, Var _) = true berghofe@10413: | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) berghofe@10413: | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) berghofe@10413: | vperm (t, u) = (t = u); berghofe@10413: berghofe@10413: fun var_perm (t, u) = haftmann@33038: vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []); berghofe@10413: wenzelm@15023: (*simple test for looping rewrite rules and stupid orientations*) wenzelm@18208: fun default_reorient thy prems lhs rhs = wenzelm@15023: rewrite_rule_extra_vars prems lhs rhs wenzelm@15023: orelse wenzelm@15023: is_Var (head_of lhs) wenzelm@15023: orelse nipkow@16305: (* turns t = x around, which causes a headache if x is a local variable - nipkow@16305: usually it is very useful :-( nipkow@16305: is_Free rhs andalso not(is_Free lhs) andalso not(Logic.occs(rhs,lhs)) nipkow@16305: andalso not(exists_subterm is_Var lhs) nipkow@16305: orelse nipkow@16305: *) wenzelm@16842: exists (fn t => Logic.occs (lhs, t)) (rhs :: prems) wenzelm@15023: orelse wenzelm@17203: null prems andalso Pattern.matches thy (lhs, rhs) berghofe@10413: (*the condition "null prems" is necessary because conditional rewrites berghofe@10413: with extra variables in the conditions may terminate although wenzelm@15023: the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*) wenzelm@15023: orelse wenzelm@15023: is_Const lhs andalso not (is_Const rhs); berghofe@10413: berghofe@10413: fun decomp_simp thm = wenzelm@15023: let wenzelm@26626: val thy = Thm.theory_of_thm thm; wenzelm@26626: val prop = Thm.prop_of thm; wenzelm@15023: val prems = Logic.strip_imp_prems prop; wenzelm@15023: val concl = Drule.strip_imp_concl (Thm.cprop_of thm); wenzelm@22902: val (lhs, rhs) = Thm.dest_equals concl handle TERM _ => wenzelm@15023: raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm); wenzelm@20579: val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs)); wenzelm@16665: val elhs = if term_of elhs aconv term_of lhs then lhs else elhs; (*share identical copies*) wenzelm@18929: val erhs = Envir.eta_contract (term_of rhs); wenzelm@15023: val perm = wenzelm@15023: var_perm (term_of elhs, erhs) andalso wenzelm@15023: not (term_of elhs aconv erhs) andalso wenzelm@15023: not (is_Var (term_of elhs)); wenzelm@16458: in (thy, prems, term_of lhs, elhs, term_of rhs, perm) end; berghofe@10413: wenzelm@12783: fun decomp_simp' thm = wenzelm@12979: let val (_, _, lhs, _, rhs, _) = decomp_simp thm in wenzelm@12783: if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", thm) wenzelm@12979: else (lhs, rhs) wenzelm@12783: end; wenzelm@12783: wenzelm@15023: fun mk_eq_True (Simpset (_, {mk_rews = {mk_eq_True, ...}, ...})) (thm, name) = wenzelm@15023: (case mk_eq_True thm of skalberg@15531: NONE => [] skalberg@15531: | SOME eq_True => wenzelm@20546: let wenzelm@20546: val (_, _, lhs, elhs, _, _) = decomp_simp eq_True; wenzelm@15023: in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end); berghofe@10413: wenzelm@15023: (*create the rewrite rule and possibly also the eq_True variant, wenzelm@15023: in case there are extra vars on the rhs*) wenzelm@15023: fun rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm2) = wenzelm@15023: let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in wenzelm@20546: if rewrite_rule_extra_vars [] lhs rhs then wenzelm@20546: mk_eq_True ss (thm2, name) @ [rrule] wenzelm@20546: else [rrule] berghofe@10413: end; berghofe@10413: wenzelm@15023: fun mk_rrule ss (thm, name) = wenzelm@15023: let val (_, prems, lhs, elhs, rhs, perm) = decomp_simp thm in wenzelm@15023: if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] wenzelm@15023: else wenzelm@15023: (*weak test for loops*) wenzelm@15023: if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs) wenzelm@15023: then mk_eq_True ss (thm, name) wenzelm@15023: else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) berghofe@10413: end; berghofe@10413: wenzelm@15023: fun orient_rrule ss (thm, name) = wenzelm@18208: let wenzelm@18208: val (thy, prems, lhs, elhs, rhs, perm) = decomp_simp thm; wenzelm@18208: val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = ss; wenzelm@18208: in wenzelm@15023: if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] wenzelm@16458: else if reorient thy prems lhs rhs then wenzelm@16458: if reorient thy prems rhs lhs wenzelm@15023: then mk_eq_True ss (thm, name) wenzelm@15023: else wenzelm@18208: (case mk_sym thm of wenzelm@18208: NONE => [] wenzelm@18208: | SOME thm' => wenzelm@18208: let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm' wenzelm@18208: in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end) wenzelm@15023: else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) berghofe@10413: end; berghofe@10413: nipkow@15199: fun extract_rews (Simpset (_, {mk_rews = {mk, ...}, ...}), thms) = wenzelm@27865: maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk thm)) thms; berghofe@10413: wenzelm@15023: fun extract_safe_rrules (ss, thm) = wenzelm@19482: maps (orient_rrule ss) (extract_rews (ss, [thm])); berghofe@10413: berghofe@10413: wenzelm@20028: (* add/del rules explicitly *) berghofe@10413: wenzelm@20028: fun comb_simps comb mk_rrule (ss, thms) = wenzelm@20028: let wenzelm@20028: val rews = extract_rews (ss, thms); wenzelm@20028: in fold (fold comb o mk_rrule) rews ss end; berghofe@10413: wenzelm@20028: fun ss addsimps thms = wenzelm@20028: comb_simps insert_rrule (mk_rrule ss) (ss, thms); berghofe@10413: wenzelm@15023: fun ss delsimps thms = wenzelm@20028: comb_simps del_rrule (map mk_rrule2 o mk_rrule ss) (ss, thms); wenzelm@15023: haftmann@27558: fun add_simp thm ss = ss addsimps [thm]; haftmann@27558: fun del_simp thm ss = ss delsimps [thm]; wenzelm@15023: wenzelm@30318: wenzelm@15023: (* congs *) berghofe@10413: skalberg@15531: fun cong_name (Const (a, _)) = SOME a skalberg@15531: | cong_name (Free (a, _)) = SOME ("Free: " ^ a) skalberg@15531: | cong_name _ = NONE; ballarin@13835: wenzelm@15023: local wenzelm@15023: wenzelm@15023: fun is_full_cong_prems [] [] = true wenzelm@15023: | is_full_cong_prems [] _ = false wenzelm@15023: | is_full_cong_prems (p :: prems) varpairs = wenzelm@15023: (case Logic.strip_assums_concl p of wenzelm@15023: Const ("==", _) $ lhs $ rhs => wenzelm@15023: let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in wenzelm@15023: is_Var x andalso forall is_Bound xs andalso haftmann@20972: not (has_duplicates (op =) xs) andalso xs = ys andalso wenzelm@20671: member (op =) varpairs (x, y) andalso wenzelm@19303: is_full_cong_prems prems (remove (op =) (x, y) varpairs) wenzelm@15023: end wenzelm@15023: | _ => false); wenzelm@15023: wenzelm@15023: fun is_full_cong thm = berghofe@10413: let wenzelm@15023: val prems = prems_of thm and concl = concl_of thm; wenzelm@15023: val (lhs, rhs) = Logic.dest_equals concl; wenzelm@15023: val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs; berghofe@10413: in haftmann@20972: f = g andalso not (has_duplicates (op =) (xs @ ys)) andalso length xs = length ys andalso wenzelm@15023: is_full_cong_prems prems (xs ~~ ys) berghofe@10413: end; berghofe@10413: wenzelm@15023: fun add_cong (ss, thm) = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: let wenzelm@22902: val (lhs, _) = Thm.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm)) wenzelm@15023: handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm); wenzelm@18929: (*val lhs = Envir.eta_contract lhs;*) wenzelm@20057: val a = the (cong_name (head_of (term_of lhs))) handle Option.Option => wenzelm@15023: raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm); haftmann@22221: val (xs, weak) = congs; haftmann@22221: val _ = if AList.defined (op =) xs a haftmann@22221: then warning ("Overwriting congruence rule for " ^ quote a) haftmann@22221: else (); krauss@30908: val xs' = AList.update (op =) (a, thm) xs; haftmann@22221: val weak' = if is_full_cong thm then weak else a :: weak; haftmann@22221: in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); berghofe@10413: wenzelm@15023: fun del_cong (ss, thm) = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: let wenzelm@15023: val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => wenzelm@15023: raise SIMPLIFIER ("Congruence not a meta-equality", thm); wenzelm@18929: (*val lhs = Envir.eta_contract lhs;*) wenzelm@20057: val a = the (cong_name (head_of lhs)) handle Option.Option => wenzelm@15023: raise SIMPLIFIER ("Congruence must start with a constant", thm); haftmann@22221: val (xs, _) = congs; haftmann@22221: val xs' = filter_out (fn (x : string, _) => x = a) xs; krauss@30908: val weak' = xs' |> map_filter (fn (a, thm) => skalberg@15531: if is_full_cong thm then NONE else SOME a); haftmann@22221: in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); berghofe@10413: wenzelm@15023: fun mk_cong (Simpset (_, {mk_rews = {mk_cong = f, ...}, ...})) = f; wenzelm@15023: wenzelm@15023: in wenzelm@15023: skalberg@15570: val (op addeqcongs) = Library.foldl add_cong; skalberg@15570: val (op deleqcongs) = Library.foldl del_cong; wenzelm@15023: wenzelm@15023: fun ss addcongs congs = ss addeqcongs map (mk_cong ss) congs; wenzelm@15023: fun ss delcongs congs = ss deleqcongs map (mk_cong ss) congs; wenzelm@15023: wenzelm@15023: end; berghofe@10413: berghofe@10413: wenzelm@15023: (* simprocs *) wenzelm@15023: wenzelm@22234: datatype simproc = wenzelm@22234: Simproc of wenzelm@22234: {name: string, wenzelm@22234: lhss: cterm list, wenzelm@22234: proc: morphism -> simpset -> cterm -> thm option, wenzelm@22234: id: stamp * thm list}; wenzelm@22234: wenzelm@22234: fun eq_simproc (Simproc {id = id1, ...}, Simproc {id = id2, ...}) = eq_procid (id1, id2); wenzelm@22008: wenzelm@22234: fun morph_simproc phi (Simproc {name, lhss, proc, id = (s, ths)}) = wenzelm@22234: Simproc wenzelm@22234: {name = name, wenzelm@22234: lhss = map (Morphism.cterm phi) lhss, wenzelm@22669: proc = Morphism.transform phi proc, wenzelm@22234: id = (s, Morphism.fact phi ths)}; wenzelm@22234: wenzelm@22234: fun make_simproc {name, lhss, proc, identifier} = wenzelm@22234: Simproc {name = name, lhss = lhss, proc = proc, id = (stamp (), identifier)}; wenzelm@22008: wenzelm@22008: fun mk_simproc name lhss proc = wenzelm@22234: make_simproc {name = name, lhss = lhss, proc = fn _ => fn ss => fn ct => wenzelm@22234: proc (ProofContext.theory_of (the_context ss)) ss (Thm.term_of ct), identifier = []}; wenzelm@22008: wenzelm@22008: (* FIXME avoid global thy and Logic.varify *) wenzelm@22008: fun simproc_i thy name = mk_simproc name o map (Thm.cterm_of thy o Logic.varify); wenzelm@24707: fun simproc thy name = simproc_i thy name o map (Syntax.read_term_global thy); wenzelm@22008: wenzelm@22008: wenzelm@15023: local berghofe@10413: wenzelm@16985: fun add_proc (proc as Proc {name, lhs, ...}) ss = wenzelm@22254: (trace_cterm false (fn () => "Adding simplification procedure " ^ quote name ^ " for") ss lhs; wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@16807: (congs, Net.insert_term eq_proc (term_of lhs, proc) procs, wenzelm@15023: mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss wenzelm@15023: handle Net.INSERT => wenzelm@15023: (warning ("Ignoring duplicate simplification procedure " ^ quote name); ss)); berghofe@10413: wenzelm@16985: fun del_proc (proc as Proc {name, lhs, ...}) ss = wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@16807: (congs, Net.delete_term eq_proc (term_of lhs, proc) procs, wenzelm@15023: mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss wenzelm@15023: handle Net.DELETE => wenzelm@15023: (warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss); berghofe@10413: wenzelm@22234: fun prep_procs (Simproc {name, lhss, proc, id}) = wenzelm@22669: lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, id = id}); wenzelm@22234: wenzelm@15023: in berghofe@10413: wenzelm@22234: fun ss addsimprocs ps = fold (fold add_proc o prep_procs) ps ss; wenzelm@22234: fun ss delsimprocs ps = fold (fold del_proc o prep_procs) ps ss; berghofe@10413: wenzelm@15023: end; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_rews *) berghofe@10413: wenzelm@15023: local wenzelm@15023: wenzelm@18208: fun map_mk_rews f = map_simpset2 (fn (congs, procs, {mk, mk_cong, mk_sym, mk_eq_True, reorient}, wenzelm@15023: termless, subgoal_tac, loop_tacs, solvers) => wenzelm@18208: let wenzelm@18208: val (mk', mk_cong', mk_sym', mk_eq_True', reorient') = wenzelm@18208: f (mk, mk_cong, mk_sym, mk_eq_True, reorient); wenzelm@18208: val mk_rews' = {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True', wenzelm@18208: reorient = reorient'}; wenzelm@18208: in (congs, procs, mk_rews', termless, subgoal_tac, loop_tacs, solvers) end); wenzelm@15023: wenzelm@15023: in berghofe@10413: wenzelm@30336: fun mksimps (Simpset (_, {mk_rews = {mk, ...}, ...})) = mk; wenzelm@30318: wenzelm@18208: fun ss setmksimps mk = ss |> map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) => wenzelm@18208: (mk, mk_cong, mk_sym, mk_eq_True, reorient)); wenzelm@15023: wenzelm@18208: fun ss setmkcong mk_cong = ss |> map_mk_rews (fn (mk, _, mk_sym, mk_eq_True, reorient) => wenzelm@18208: (mk, mk_cong, mk_sym, mk_eq_True, reorient)); berghofe@10413: wenzelm@18208: fun ss setmksym mk_sym = ss |> map_mk_rews (fn (mk, mk_cong, _, mk_eq_True, reorient) => wenzelm@18208: (mk, mk_cong, mk_sym, mk_eq_True, reorient)); berghofe@10413: wenzelm@18208: fun ss setmkeqTrue mk_eq_True = ss |> map_mk_rews (fn (mk, mk_cong, mk_sym, _, reorient) => wenzelm@18208: (mk, mk_cong, mk_sym, mk_eq_True, reorient)); wenzelm@18208: wenzelm@18208: fun set_reorient reorient = map_mk_rews (fn (mk, mk_cong, mk_sym, mk_eq_True, _) => wenzelm@18208: (mk, mk_cong, mk_sym, mk_eq_True, reorient)); wenzelm@15023: wenzelm@15023: end; wenzelm@15023: skalberg@14242: berghofe@10413: (* termless *) berghofe@10413: wenzelm@15023: fun ss settermless termless = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); skalberg@15006: skalberg@15006: wenzelm@15023: (* tactics *) skalberg@15006: wenzelm@15023: fun ss setsubgoaler subgoal_tac = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) => wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); skalberg@15006: wenzelm@17882: fun ss setloop' tac = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) => wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers)); skalberg@15006: wenzelm@17882: fun ss setloop tac = ss setloop' (K tac); wenzelm@17882: wenzelm@17882: fun ss addloop' (name, tac) = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, haftmann@21286: (if AList.defined (op =) loop_tacs name haftmann@21286: then warning ("Overwriting looper " ^ quote name) haftmann@21286: else (); AList.update (op =) (name, tac) loop_tacs), wenzelm@15023: solvers)); skalberg@15006: wenzelm@17882: fun ss addloop (name, tac) = ss addloop' (name, K tac); wenzelm@17882: wenzelm@15023: fun ss delloop name = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => haftmann@21286: (congs, procs, mk_rews, termless, subgoal_tac, haftmann@21286: (if AList.defined (op =) loop_tacs name haftmann@21286: then () haftmann@21286: else warning ("No such looper in simpset: " ^ quote name); haftmann@21286: AList.delete (op =) name loop_tacs), solvers)); skalberg@15006: wenzelm@15023: fun ss setSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, (unsafe_solvers, _)) => wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver]))); skalberg@15006: wenzelm@15023: fun ss addSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, haftmann@22717: subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers))); skalberg@15006: wenzelm@15023: fun ss setSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, ([solver], solvers))); skalberg@15006: wenzelm@15023: fun ss addSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, haftmann@22717: subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers))); skalberg@15006: wenzelm@15023: fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, termless, wenzelm@15023: subgoal_tac, loop_tacs, (solvers, solvers))); skalberg@15006: skalberg@15006: wenzelm@18208: (* empty *) wenzelm@18208: wenzelm@18208: fun init_ss mk_rews termless subgoal_tac solvers = wenzelm@32738: make_simpset ((Net.empty, [], (0, []), (0, Unsynchronized.ref false), NONE), wenzelm@18208: (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers)); wenzelm@18208: wenzelm@18208: fun clear_ss (ss as Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...})) = wenzelm@18208: init_ss mk_rews termless subgoal_tac solvers wenzelm@18208: |> inherit_context ss; wenzelm@18208: wenzelm@18208: val basic_mk_rews: mk_rews = wenzelm@18208: {mk = fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], wenzelm@18208: mk_cong = I, wenzelm@18208: mk_sym = SOME o Drule.symmetric_fun, wenzelm@18208: mk_eq_True = K NONE, wenzelm@18208: reorient = default_reorient}; wenzelm@18208: wenzelm@29269: val empty_ss = init_ss basic_mk_rews TermOrd.termless (K (K no_tac)) ([], []); wenzelm@18208: wenzelm@18208: wenzelm@18208: (* merge *) (*NOTE: ignores some fields of 2nd simpset*) wenzelm@18208: wenzelm@18208: fun merge_ss (ss1, ss2) = wenzelm@24358: if pointer_eq (ss1, ss2) then ss1 wenzelm@24358: else wenzelm@24358: let wenzelm@24358: val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth = depth1, context = _}, wenzelm@24358: {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac, wenzelm@24358: loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; wenzelm@24358: val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = depth2, context = _}, wenzelm@24358: {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _, wenzelm@24358: loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; wenzelm@30356: wenzelm@24358: val rules' = Net.merge eq_rrule (rules1, rules2); wenzelm@33520: val prems' = Thm.merge_thms (prems1, prems2); wenzelm@24358: val bounds' = if #1 bounds1 < #1 bounds2 then bounds2 else bounds1; wenzelm@24358: val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1; wenzelm@31298: val congs' = merge (Thm.eq_thm_prop o pairself #2) (congs1, congs2); wenzelm@24358: val weak' = merge (op =) (weak1, weak2); wenzelm@24358: val procs' = Net.merge eq_proc (procs1, procs2); wenzelm@24358: val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2); wenzelm@24358: val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2); wenzelm@24358: val solvers' = merge eq_solver (solvers1, solvers2); wenzelm@24358: in wenzelm@24358: make_simpset ((rules', prems', bounds', depth', NONE), ((congs', weak'), procs', wenzelm@24358: mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) wenzelm@24358: end; wenzelm@18208: wenzelm@18208: wenzelm@30356: (* dest_ss *) wenzelm@30356: wenzelm@30356: fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) = wenzelm@30356: {simps = Net.entries rules wenzelm@30356: |> map (fn {name, thm, ...} => (name, thm)), wenzelm@30356: procs = Net.entries procs wenzelm@30356: |> map (fn Proc {name, lhs, id, ...} => ((name, lhs), id)) wenzelm@30356: |> partition_eq (eq_snd eq_procid) wenzelm@30356: |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)), krauss@30908: congs = #1 congs, wenzelm@30356: weak_congs = #2 congs, wenzelm@30356: loopers = map fst loop_tacs, wenzelm@30356: unsafe_solvers = map solver_name (#1 solvers), wenzelm@30356: safe_solvers = map solver_name (#2 solvers)}; wenzelm@30356: wenzelm@30356: skalberg@15006: berghofe@10413: (** rewriting **) berghofe@10413: berghofe@10413: (* berghofe@10413: Uses conversions, see: berghofe@10413: L C Paulson, A higher-order implementation of rewriting, berghofe@10413: Science of Computer Programming 3 (1983), pages 119-149. berghofe@10413: *) berghofe@10413: wenzelm@16985: fun check_conv msg ss thm thm' = berghofe@10413: let berghofe@25472: val thm'' = transitive thm thm' handle THM _ => berghofe@25472: transitive thm (transitive berghofe@25472: (symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm') wenzelm@22254: in if msg then trace_thm (fn () => "SUCCEEDED") ss thm' else (); SOME thm'' end berghofe@10413: handle THM _ => wenzelm@26626: let wenzelm@26626: val thy = Thm.theory_of_thm thm; wenzelm@26626: val _ $ _ $ prop0 = Thm.prop_of thm; wenzelm@26626: in wenzelm@22254: trace_thm (fn () => "Proved wrong thm (Check subgoaler?)") ss thm'; wenzelm@22254: trace_term false (fn () => "Should have proved:") ss thy prop0; skalberg@15531: NONE berghofe@10413: end; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_procrule *) berghofe@10413: wenzelm@16985: fun mk_procrule ss thm = wenzelm@15023: let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in wenzelm@15023: if rewrite_rule_extra_vars prems lhs rhs wenzelm@16985: then (warn_thm "Extra vars on rhs:" ss thm; []) wenzelm@15023: else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}] berghofe@10413: end; berghofe@10413: berghofe@10413: wenzelm@15023: (* rewritec: conversion to apply the meta simpset to a term *) berghofe@10413: wenzelm@15023: (*Since the rewriting strategy is bottom-up, we avoid re-normalizing already wenzelm@15023: normalized terms by carrying around the rhs of the rewrite rule just wenzelm@15023: applied. This is called the `skeleton'. It is decomposed in parallel wenzelm@15023: with the term. Once a Var is encountered, the corresponding term is wenzelm@15023: already in normal form. wenzelm@15023: skel0 is a dummy skeleton that is to enforce complete normalization.*) wenzelm@15023: berghofe@10413: val skel0 = Bound 0; berghofe@10413: wenzelm@15023: (*Use rhs as skeleton only if the lhs does not contain unnormalized bits. wenzelm@15023: The latter may happen iff there are weak congruence rules for constants wenzelm@15023: in the lhs.*) berghofe@10413: wenzelm@15023: fun uncond_skel ((_, weak), (lhs, rhs)) = wenzelm@15023: if null weak then rhs (*optimization*) wenzelm@20671: else if exists_Const (member (op =) weak o #1) lhs then skel0 wenzelm@15023: else rhs; wenzelm@15023: wenzelm@15023: (*Behaves like unconditional rule if rhs does not contain vars not in the lhs. wenzelm@15023: Otherwise those vars may become instantiated with unnormalized terms wenzelm@15023: while the premises are solved.*) wenzelm@15023: wenzelm@32797: fun cond_skel (args as (_, (lhs, rhs))) = haftmann@33038: if subset (op =) (Term.add_vars rhs [], Term.add_vars lhs []) then uncond_skel args berghofe@10413: else skel0; berghofe@10413: berghofe@10413: (* wenzelm@15023: Rewriting -- we try in order: berghofe@10413: (1) beta reduction berghofe@10413: (2) unconditional rewrite rules berghofe@10413: (3) conditional rewrite rules berghofe@10413: (4) simplification procedures berghofe@10413: berghofe@10413: IMPORTANT: rewrite rules must not introduce new Vars or TVars! berghofe@10413: *) berghofe@10413: wenzelm@16458: fun rewritec (prover, thyt, maxt) ss t = berghofe@10413: let wenzelm@24124: val ctxt = the_context ss; wenzelm@15023: val Simpset ({rules, ...}, {congs, procs, termless, ...}) = ss; berghofe@10413: val eta_thm = Thm.eta_conversion t; wenzelm@22902: val eta_t' = Thm.rhs_of eta_thm; berghofe@10413: val eta_t = term_of eta_t'; wenzelm@20546: fun rew {thm, name, lhs, elhs, extra, fo, perm} = berghofe@10413: let wenzelm@32797: val prop = Thm.prop_of thm; wenzelm@20546: val (rthm, elhs') = wenzelm@20546: if maxt = ~1 orelse not extra then (thm, elhs) wenzelm@22902: else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs); wenzelm@22902: val insts = wenzelm@22902: if fo then Thm.first_order_match (elhs', eta_t') wenzelm@22902: else Thm.match (elhs', eta_t'); berghofe@10413: val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm); wenzelm@14643: val prop' = Thm.prop_of thm'; wenzelm@21576: val unconditional = (Logic.count_prems prop' = 0); berghofe@10413: val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop') berghofe@10413: in nipkow@11295: if perm andalso not (termless (rhs', lhs')) wenzelm@22254: then (trace_named_thm (fn () => "Cannot apply permutative rewrite rule") ss (thm, name); wenzelm@22254: trace_thm (fn () => "Term does not become smaller:") ss thm'; NONE) wenzelm@22254: else (trace_named_thm (fn () => "Applying instance of rewrite rule") ss (thm, name); berghofe@10413: if unconditional berghofe@10413: then wenzelm@22254: (trace_thm (fn () => "Rewriting:") ss thm'; berghofe@10413: let val lr = Logic.dest_equals prop; wenzelm@16985: val SOME thm'' = check_conv false ss eta_thm thm' skalberg@15531: in SOME (thm'', uncond_skel (congs, lr)) end) berghofe@10413: else wenzelm@22254: (trace_thm (fn () => "Trying to rewrite:") ss thm'; wenzelm@24124: if simp_depth ss > Config.get ctxt simp_depth_limit nipkow@16042: then let val s = "simp_depth_limit exceeded - giving up" wenzelm@22892: in trace false (fn () => s) ss; warning s; NONE end nipkow@16042: else nipkow@16042: case prover ss thm' of wenzelm@22254: NONE => (trace_thm (fn () => "FAILED") ss thm'; NONE) skalberg@15531: | SOME thm2 => wenzelm@16985: (case check_conv true ss eta_thm thm2 of skalberg@15531: NONE => NONE | skalberg@15531: SOME thm2' => berghofe@10413: let val concl = Logic.strip_imp_concl prop berghofe@10413: val lr = Logic.dest_equals concl nipkow@16042: in SOME (thm2', cond_skel (congs, lr)) end))) berghofe@10413: end berghofe@10413: skalberg@15531: fun rews [] = NONE berghofe@10413: | rews (rrule :: rrules) = skalberg@15531: let val opt = rew rrule handle Pattern.MATCH => NONE skalberg@15531: in case opt of NONE => rews rrules | some => some end; berghofe@10413: berghofe@10413: fun sort_rrules rrs = let wenzelm@14643: fun is_simple({thm, ...}:rrule) = case Thm.prop_of thm of berghofe@10413: Const("==",_) $ _ $ _ => true wenzelm@12603: | _ => false berghofe@10413: fun sort [] (re1,re2) = re1 @ re2 wenzelm@12603: | sort (rr::rrs) (re1,re2) = if is_simple rr berghofe@10413: then sort rrs (rr::re1,re2) berghofe@10413: else sort rrs (re1,rr::re2) berghofe@10413: in sort rrs ([],[]) end berghofe@10413: skalberg@15531: fun proc_rews [] = NONE wenzelm@15023: | proc_rews (Proc {name, proc, lhs, ...} :: ps) = wenzelm@17203: if Pattern.matches thyt (Thm.term_of lhs, Thm.term_of t) then wenzelm@22254: (debug_term false (fn () => "Trying procedure " ^ quote name ^ " on:") ss thyt eta_t; wenzelm@23938: case proc ss eta_t' of wenzelm@22892: NONE => (debug false (fn () => "FAILED") ss; proc_rews ps) skalberg@15531: | SOME raw_thm => wenzelm@22254: (trace_thm (fn () => "Procedure " ^ quote name ^ " produced rewrite rule:") wenzelm@22254: ss raw_thm; wenzelm@16985: (case rews (mk_procrule ss raw_thm) of wenzelm@22254: NONE => (trace_cterm true (fn () => "IGNORED result of simproc " ^ quote name ^ wenzelm@16985: " -- does not match") ss t; proc_rews ps) berghofe@10413: | some => some))) berghofe@10413: else proc_rews ps; berghofe@10413: in case eta_t of skalberg@15531: Abs _ $ _ => SOME (transitive eta_thm berghofe@12155: (beta_conversion false eta_t'), skel0) berghofe@10413: | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of skalberg@15531: NONE => proc_rews (Net.match_term procs eta_t) berghofe@10413: | some => some) berghofe@10413: end; berghofe@10413: berghofe@10413: berghofe@10413: (* conversion to apply a congruence rule to a term *) berghofe@10413: krauss@30908: fun congc prover ss maxt cong t = wenzelm@22902: let val rthm = Thm.incr_indexes (maxt + 1) cong; wenzelm@22902: val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm))); wenzelm@22902: val insts = Thm.match (rlhs, t) wenzelm@22902: (* Thm.match can raise Pattern.MATCH; berghofe@10413: is handled when congc is called *) berghofe@10413: val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm); wenzelm@32797: val _ = trace_thm (fn () => "Applying congruence rule:") ss thm'; wenzelm@22254: fun err (msg, thm) = (trace_thm (fn () => msg) ss thm; NONE) berghofe@10413: in case prover thm' of skalberg@15531: NONE => err ("Congruence proof failed. Could not prove", thm') wenzelm@16985: | SOME thm2 => (case check_conv true ss (Drule.beta_eta_conversion t) thm2 of skalberg@15531: NONE => err ("Congruence proof failed. Should not have proved", thm2) skalberg@15531: | SOME thm2' => wenzelm@22902: if op aconv (pairself term_of (Thm.dest_equals (cprop_of thm2'))) skalberg@15531: then NONE else SOME thm2') berghofe@10413: end; berghofe@10413: berghofe@10413: val (cA, (cB, cC)) = wenzelm@22902: apsnd Thm.dest_equals (Thm.dest_implies (hd (cprems_of Drule.imp_cong))); berghofe@10413: skalberg@15531: fun transitive1 NONE NONE = NONE skalberg@15531: | transitive1 (SOME thm1) NONE = SOME thm1 skalberg@15531: | transitive1 NONE (SOME thm2) = SOME thm2 skalberg@15531: | transitive1 (SOME thm1) (SOME thm2) = SOME (transitive thm1 thm2) berghofe@10413: skalberg@15531: fun transitive2 thm = transitive1 (SOME thm); skalberg@15531: fun transitive3 thm = transitive1 thm o SOME; berghofe@13607: wenzelm@16458: fun bottomc ((simprem, useprem, mutsimp), prover, thy, maxidx) = berghofe@10413: let wenzelm@15023: fun botc skel ss t = skalberg@15531: if is_Var skel then NONE berghofe@10413: else wenzelm@15023: (case subc skel ss t of skalberg@15531: some as SOME thm1 => wenzelm@22902: (case rewritec (prover, thy, maxidx) ss (Thm.rhs_of thm1) of skalberg@15531: SOME (thm2, skel2) => berghofe@13607: transitive2 (transitive thm1 thm2) wenzelm@22902: (botc skel2 ss (Thm.rhs_of thm2)) skalberg@15531: | NONE => some) skalberg@15531: | NONE => wenzelm@16458: (case rewritec (prover, thy, maxidx) ss t of skalberg@15531: SOME (thm2, skel2) => transitive2 thm2 wenzelm@22902: (botc skel2 ss (Thm.rhs_of thm2)) skalberg@15531: | NONE => NONE)) berghofe@10413: wenzelm@15023: and try_botc ss t = wenzelm@15023: (case botc skel0 ss t of skalberg@15531: SOME trec1 => trec1 | NONE => (reflexive t)) berghofe@10413: wenzelm@15023: and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 = berghofe@10413: (case term_of t0 of wenzelm@32797: Abs (a, T, _) => wenzelm@15023: let wenzelm@20079: val b = Name.bound (#1 bounds); wenzelm@16985: val (v, t') = Thm.dest_abs (SOME b) t0; wenzelm@16985: val b' = #1 (Term.dest_Free (Thm.term_of v)); wenzelm@21962: val _ = wenzelm@21962: if b <> b' then wenzelm@33379: warning ("Simplifier: renamed bound variable " ^ quote b ^ " to " ^ quote b') wenzelm@21962: else (); wenzelm@17614: val ss' = add_bound ((b', T), a) ss; wenzelm@15023: val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0; wenzelm@15023: in case botc skel' ss' t' of skalberg@15531: SOME thm => SOME (abstract_rule a v thm) skalberg@15531: | NONE => NONE berghofe@10413: end berghofe@10413: | t $ _ => (case t of wenzelm@15023: Const ("==>", _) $ _ => impc t0 ss berghofe@10413: | Abs _ => berghofe@10413: let val thm = beta_conversion false t0 wenzelm@22902: in case subc skel0 ss (Thm.rhs_of thm) of skalberg@15531: NONE => SOME thm skalberg@15531: | SOME thm' => SOME (transitive thm thm') berghofe@10413: end berghofe@10413: | _ => berghofe@10413: let fun appc () = berghofe@10413: let berghofe@10413: val (tskel, uskel) = case skel of berghofe@10413: tskel $ uskel => (tskel, uskel) berghofe@10413: | _ => (skel0, skel0); wenzelm@10767: val (ct, cu) = Thm.dest_comb t0 berghofe@10413: in wenzelm@15023: (case botc tskel ss ct of skalberg@15531: SOME thm1 => wenzelm@15023: (case botc uskel ss cu of skalberg@15531: SOME thm2 => SOME (combination thm1 thm2) skalberg@15531: | NONE => SOME (combination thm1 (reflexive cu))) skalberg@15531: | NONE => wenzelm@15023: (case botc uskel ss cu of skalberg@15531: SOME thm1 => SOME (combination (reflexive ct) thm1) skalberg@15531: | NONE => NONE)) berghofe@10413: end berghofe@10413: val (h, ts) = strip_comb t ballarin@13835: in case cong_name h of skalberg@15531: SOME a => haftmann@17232: (case AList.lookup (op =) (fst congs) a of skalberg@15531: NONE => appc () skalberg@15531: | SOME cong => wenzelm@15023: (*post processing: some partial applications h t1 ... tj, j <= length ts, wenzelm@15023: may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*) berghofe@10413: (let wenzelm@16985: val thm = congc (prover ss) ss maxidx cong t0; wenzelm@22902: val t = the_default t0 (Option.map Thm.rhs_of thm); wenzelm@10767: val (cl, cr) = Thm.dest_comb t berghofe@10413: val dVar = Var(("", 0), dummyT) berghofe@10413: val skel = berghofe@10413: list_comb (h, replicate (length ts) dVar) wenzelm@15023: in case botc skel ss cl of skalberg@15531: NONE => thm skalberg@15531: | SOME thm' => transitive3 thm berghofe@12155: (combination thm' (reflexive cr)) wenzelm@20057: end handle Pattern.MATCH => appc ())) berghofe@10413: | _ => appc () berghofe@10413: end) skalberg@15531: | _ => NONE) berghofe@10413: wenzelm@15023: and impc ct ss = wenzelm@15023: if mutsimp then mut_impc0 [] ct [] [] ss else nonmut_impc ct ss berghofe@10413: wenzelm@15023: and rules_of_prem ss prem = berghofe@13607: if maxidx_of_term (term_of prem) <> ~1 berghofe@13607: then (trace_cterm true wenzelm@22254: (fn () => "Cannot add premise as rewrite rule because it contains (type) unknowns:") wenzelm@22254: ss prem; ([], NONE)) berghofe@13607: else berghofe@13607: let val asm = assume prem skalberg@15531: in (extract_safe_rrules (ss, asm), SOME asm) end berghofe@10413: wenzelm@15023: and add_rrules (rrss, asms) ss = wenzelm@20028: (fold o fold) insert_rrule rrss ss |> add_prems (map_filter I asms) berghofe@10413: wenzelm@23178: and disch r prem eq = berghofe@13607: let wenzelm@22902: val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq); berghofe@13607: val eq' = implies_elim (Thm.instantiate berghofe@13607: ([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong) berghofe@13607: (implies_intr prem eq) berghofe@13607: in if not r then eq' else berghofe@10413: let wenzelm@22902: val (prem', concl) = Thm.dest_implies lhs; wenzelm@22902: val (prem'', _) = Thm.dest_implies rhs berghofe@13607: in transitive (transitive berghofe@13607: (Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)]) berghofe@13607: Drule.swap_prems_eq) eq') berghofe@13607: (Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)]) berghofe@13607: Drule.swap_prems_eq) berghofe@10413: end berghofe@10413: end berghofe@10413: berghofe@13607: and rebuild [] _ _ _ _ eq = eq wenzelm@32797: | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ss eq = berghofe@13607: let wenzelm@15023: val ss' = add_rrules (rev rrss, rev asms) ss; berghofe@13607: val concl' = wenzelm@22902: Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq)); wenzelm@23178: val dprem = Option.map (disch false prem) wenzelm@16458: in case rewritec (prover, thy, maxidx) ss' concl' of skalberg@15531: NONE => rebuild prems concl' rrss asms ss (dprem eq) wenzelm@23178: | SOME (eq', _) => transitive2 (fold (disch false) wenzelm@23178: prems (the (transitive3 (dprem eq) eq'))) wenzelm@22902: (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ss) berghofe@13607: end wenzelm@15023: wenzelm@15023: and mut_impc0 prems concl rrss asms ss = berghofe@13607: let berghofe@13607: val prems' = strip_imp_prems concl; wenzelm@15023: val (rrss', asms') = split_list (map (rules_of_prem ss) prems') berghofe@13607: in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') wenzelm@15023: (asms @ asms') [] [] [] [] ss ~1 ~1 berghofe@13607: end wenzelm@15023: wenzelm@15023: and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k = wenzelm@33245: transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1 wenzelm@33245: (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE) berghofe@13607: (if changed > 0 then berghofe@13607: mut_impc (rev prems') concl (rev rrss') (rev asms') wenzelm@15023: [] [] [] [] ss ~1 changed wenzelm@15023: else rebuild prems' concl rrss' asms' ss wenzelm@15023: (botc skel0 (add_rrules (rev rrss', rev asms') ss) concl)) berghofe@13607: berghofe@13607: | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) wenzelm@15023: prems' rrss' asms' eqns ss changed k = skalberg@15531: case (if k = 0 then NONE else botc skel0 (add_rrules wenzelm@15023: (rev rrss' @ rrss, rev asms' @ asms) ss) prem) of skalberg@15531: NONE => mut_impc prems concl rrss asms (prem :: prems') skalberg@15531: (rrs :: rrss') (asm :: asms') (NONE :: eqns) ss changed berghofe@13607: (if k = 0 then 0 else k - 1) skalberg@15531: | SOME eqn => berghofe@13607: let wenzelm@22902: val prem' = Thm.rhs_of eqn; berghofe@13607: val tprems = map term_of prems; wenzelm@33029: val i = 1 + fold Integer.max (map (fn p => wenzelm@33029: find_index (fn q => q aconv p) tprems) (#hyps (rep_thm eqn))) ~1; wenzelm@15023: val (rrs', asm') = rules_of_prem ss prem' berghofe@13607: in mut_impc prems concl rrss asms (prem' :: prems') wenzelm@23178: (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true) haftmann@33955: ((uncurry take) (i, prems)) wenzelm@18470: (Drule.imp_cong_rule eqn (reflexive (Drule.list_implies haftmann@33955: ((uncurry drop) (i, prems), concl))))) :: eqns) wenzelm@20671: ss (length prems') ~1 berghofe@13607: end berghofe@13607: wenzelm@15023: (*legacy code - only for backwards compatibility*) wenzelm@15023: and nonmut_impc ct ss = wenzelm@22902: let val (prem, conc) = Thm.dest_implies ct; skalberg@15531: val thm1 = if simprem then botc skel0 ss prem else NONE; wenzelm@22902: val prem1 = the_default prem (Option.map Thm.rhs_of thm1); wenzelm@15023: val ss1 = if not useprem then ss else add_rrules wenzelm@15023: (apsnd single (apfst single (rules_of_prem ss prem1))) ss wenzelm@15023: in (case botc skel0 ss1 conc of skalberg@15531: NONE => (case thm1 of skalberg@15531: NONE => NONE wenzelm@18470: | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (reflexive conc))) skalberg@15531: | SOME thm2 => wenzelm@23178: let val thm2' = disch false prem1 thm2 berghofe@10413: in (case thm1 of skalberg@15531: NONE => SOME thm2' skalberg@15531: | SOME thm1' => wenzelm@18470: SOME (transitive (Drule.imp_cong_rule thm1' (reflexive conc)) thm2')) berghofe@10413: end) berghofe@10413: end berghofe@10413: wenzelm@15023: in try_botc end; berghofe@10413: berghofe@10413: wenzelm@15023: (* Meta-rewriting: rewrites t to u and returns the theorem t==u *) berghofe@10413: berghofe@10413: (* berghofe@10413: Parameters: berghofe@10413: mode = (simplify A, berghofe@10413: use A in simplifying B, berghofe@10413: use prems of B (if B is again a meta-impl.) to simplify A) berghofe@10413: when simplifying A ==> B berghofe@10413: prover: how to solve premises in conditional rewrites and congruences berghofe@10413: *) berghofe@10413: wenzelm@32738: val debug_bounds = Unsynchronized.ref false; wenzelm@17705: wenzelm@21962: fun check_bounds ss ct = wenzelm@21962: if ! debug_bounds then wenzelm@21962: let wenzelm@21962: val Simpset ({bounds = (_, bounds), ...}, _) = ss; wenzelm@21962: val bs = fold_aterms (fn Free (x, _) => wenzelm@21962: if Name.is_bound x andalso not (AList.defined eq_bound bounds x) wenzelm@21962: then insert (op =) x else I wenzelm@21962: | _ => I) (term_of ct) []; wenzelm@21962: in wenzelm@21962: if null bs then () wenzelm@22892: else print_term ss true ("Simplifier: term contains loose bounds: " ^ commas_quote bs) wenzelm@21962: (Thm.theory_of_cterm ct) (Thm.term_of ct) wenzelm@21962: end wenzelm@21962: else (); wenzelm@17614: wenzelm@19052: fun rewrite_cterm mode prover raw_ss raw_ct = wenzelm@17882: let wenzelm@26626: val thy = Thm.theory_of_cterm raw_ct; wenzelm@20260: val ct = Thm.adjust_maxidx_cterm ~1 raw_ct; wenzelm@32797: val {maxidx, ...} = Thm.rep_cterm ct; wenzelm@22892: val ss = inc_simp_depth (activate_context thy raw_ss); wenzelm@22892: val depth = simp_depth ss; wenzelm@21962: val _ = wenzelm@22892: if depth mod 20 = 0 then wenzelm@22892: warning ("Simplification depth " ^ string_of_int depth) wenzelm@21962: else (); wenzelm@22254: val _ = trace_cterm false (fn () => "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:") ss ct; wenzelm@17882: val _ = check_bounds ss ct; wenzelm@22892: in bottomc (mode, Option.map Drule.flexflex_unique oo prover, thy, maxidx) ss ct end; berghofe@10413: wenzelm@21708: val simple_prover = wenzelm@21708: SINGLE o (fn ss => ALLGOALS (resolve_tac (prems_of_ss ss))); wenzelm@21708: wenzelm@21708: fun rewrite _ [] ct = Thm.reflexive ct haftmann@27582: | rewrite full thms ct = rewrite_cterm (full, false, false) simple_prover haftmann@27582: (theory_context (Thm.theory_of_cterm ct) empty_ss addsimps thms) ct; wenzelm@11672: wenzelm@23598: fun simplify full thms = Conv.fconv_rule (rewrite full thms); wenzelm@21708: val rewrite_rule = simplify true; wenzelm@21708: wenzelm@15023: (*simple term rewriting -- no proof*) wenzelm@16458: fun rewrite_term thy rules procs = wenzelm@17203: Pattern.rewrite_term thy (map decomp_simp' rules) procs; wenzelm@15023: wenzelm@22902: fun rewrite_thm mode prover ss = Conv.fconv_rule (rewrite_cterm mode prover ss); berghofe@10413: wenzelm@23536: (*Rewrite the subgoals of a proof state (represented by a theorem)*) wenzelm@21708: fun rewrite_goals_rule thms th = wenzelm@23584: Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover wenzelm@23584: (theory_context (Thm.theory_of_thm th) empty_ss addsimps thms))) th; berghofe@10413: wenzelm@15023: (*Rewrite the subgoal of a proof state (represented by a theorem)*) skalberg@15011: fun rewrite_goal_rule mode prover ss i thm = wenzelm@23536: if 0 < i andalso i <= Thm.nprems_of thm wenzelm@23584: then Conv.gconv_rule (rewrite_cterm mode prover ss) i thm wenzelm@23536: else raise THM ("rewrite_goal_rule", i, [thm]); berghofe@10413: wenzelm@20228: wenzelm@21708: (** meta-rewriting tactics **) wenzelm@21708: wenzelm@28839: (*Rewrite all subgoals*) wenzelm@21708: fun rewrite_goals_tac defs = PRIMITIVE (rewrite_goals_rule defs); wenzelm@21708: fun rewtac def = rewrite_goals_tac [def]; wenzelm@21708: wenzelm@28839: (*Rewrite one subgoal*) wenzelm@25203: fun asm_rewrite_goal_tac mode prover_tac ss i thm = wenzelm@25203: if 0 < i andalso i <= Thm.nprems_of thm then wenzelm@25203: Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ss) i thm) wenzelm@25203: else Seq.empty; wenzelm@23536: wenzelm@23536: fun rewrite_goal_tac rews = wenzelm@23536: let val ss = empty_ss addsimps rews in wenzelm@23536: fn i => fn st => asm_rewrite_goal_tac (true, false, false) (K no_tac) wenzelm@23536: (theory_context (Thm.theory_of_thm st) ss) i st wenzelm@23536: end; wenzelm@23536: wenzelm@21708: (*Prunes all redundant parameters from the proof state by rewriting. wenzelm@21708: DOES NOT rewrite main goal, where quantification over an unused bound wenzelm@21708: variable is sometimes done to avoid the need for cut_facts_tac.*) wenzelm@21708: val prune_params_tac = rewrite_goals_tac [triv_forall_equality]; wenzelm@21708: wenzelm@21708: wenzelm@21708: (* for folding definitions, handling critical pairs *) wenzelm@21708: wenzelm@21708: (*The depth of nesting in a term*) wenzelm@32797: fun term_depth (Abs (_, _, t)) = 1 + term_depth t wenzelm@32797: | term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t) wenzelm@21708: | term_depth _ = 0; wenzelm@21708: wenzelm@21708: val lhs_of_thm = #1 o Logic.dest_equals o prop_of; wenzelm@21708: wenzelm@21708: (*folding should handle critical pairs! E.g. K == Inl(0), S == Inr(Inl(0)) wenzelm@21708: Returns longest lhs first to avoid folding its subexpressions.*) wenzelm@21708: fun sort_lhs_depths defs = wenzelm@21708: let val keylist = AList.make (term_depth o lhs_of_thm) defs wenzelm@21708: val keys = sort_distinct (rev_order o int_ord) (map #2 keylist) wenzelm@21708: in map (AList.find (op =) keylist) keys end; wenzelm@21708: wenzelm@21708: val rev_defs = sort_lhs_depths o map symmetric; wenzelm@21708: wenzelm@21708: fun fold_rule defs = fold rewrite_rule (rev_defs defs); wenzelm@21708: fun fold_goals_tac defs = EVERY (map rewrite_goals_tac (rev_defs defs)); wenzelm@21708: wenzelm@21708: wenzelm@20228: (* HHF normal form: !! before ==>, outermost !! generalized *) wenzelm@20228: wenzelm@20228: local wenzelm@20228: wenzelm@21565: fun gen_norm_hhf ss th = wenzelm@21565: (if Drule.is_norm_hhf (Thm.prop_of th) then th wenzelm@26424: else Conv.fconv_rule wenzelm@26424: (rewrite_cterm (true, false, false) (K (K NONE)) (theory_context (Thm.theory_of_thm th) ss)) th) wenzelm@21565: |> Thm.adjust_maxidx_thm ~1 wenzelm@21565: |> Drule.gen_all; wenzelm@20228: wenzelm@28620: val hhf_ss = empty_ss addsimps Drule.norm_hhf_eqs; wenzelm@20228: wenzelm@20228: in wenzelm@20228: wenzelm@26424: val norm_hhf = gen_norm_hhf hhf_ss; wenzelm@26424: val norm_hhf_protect = gen_norm_hhf (hhf_ss addeqcongs [Drule.protect_cong]); wenzelm@20228: wenzelm@20228: end; wenzelm@20228: berghofe@10413: end; berghofe@10413: wenzelm@32738: structure Basic_Meta_Simplifier: BASIC_META_SIMPLIFIER = MetaSimplifier; wenzelm@32738: open Basic_Meta_Simplifier;