berghofe@10413: (* Title: Pure/meta_simplifier.ML berghofe@10413: ID: $Id$ wenzelm@11672: Author: Tobias Nipkow and Stefan Berghofer 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@15023: setloop addloop delloop setSSolver addSSolver setSolver addSolver; skalberg@15006: wenzelm@11672: signature BASIC_META_SIMPLIFIER = wenzelm@11672: sig wenzelm@15023: val debug_simp: bool ref wenzelm@11672: val trace_simp: bool ref nipkow@13828: val simp_depth_limit: int ref wenzelm@15023: type rrule wenzelm@15023: type cong skalberg@15006: type solver skalberg@15006: val mk_solver: string -> (thm list -> int -> tactic) -> solver wenzelm@15023: type simpset wenzelm@15023: type proc skalberg@15006: val rep_ss: simpset -> wenzelm@15023: {rules: rrule Net.net, wenzelm@15023: prems: thm list, berghofe@15249: bounds: int, wenzelm@15023: depth: int} * wenzelm@15023: {congs: (string * cong) list * string list, wenzelm@15023: procs: proc Net.net, wenzelm@15023: mk_rews: wenzelm@15023: {mk: thm -> thm list, wenzelm@15023: mk_cong: thm -> thm, wenzelm@15023: mk_sym: thm -> thm option, wenzelm@15023: mk_eq_True: thm -> thm option}, wenzelm@15023: termless: term * term -> bool, skalberg@15006: subgoal_tac: simpset -> int -> tactic, skalberg@15006: loop_tacs: (string * (int -> tactic)) list, wenzelm@15023: solvers: solver list * solver list} skalberg@15006: val print_ss: simpset -> unit wenzelm@15023: val empty_ss: simpset wenzelm@15023: val merge_ss: simpset * simpset -> simpset wenzelm@15023: type simproc wenzelm@15023: val mk_simproc: string -> cterm list -> wenzelm@15023: (Sign.sg -> 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@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@15023: val setloop: 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 skalberg@15006: val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic 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@15023: val clear_ss: simpset -> simpset wenzelm@15023: exception SIMPROC_FAIL of string * exn wenzelm@15023: val simproc_i: Sign.sg -> string -> term list wenzelm@15023: -> (Sign.sg -> simpset -> term -> thm option) -> simproc wenzelm@15023: val simproc: Sign.sg -> string -> string list wenzelm@15023: -> (Sign.sg -> simpset -> term -> thm option) -> simproc wenzelm@11672: val rewrite_cterm: bool * bool * bool -> wenzelm@15023: (simpset -> thm -> thm option) -> simpset -> cterm -> thm wenzelm@15023: val rewrite_aux: (simpset -> thm -> thm option) -> bool -> thm list -> cterm -> thm wenzelm@15023: val simplify_aux: (simpset -> thm -> thm option) -> bool -> thm list -> thm -> thm berghofe@13196: val rewrite_term: Sign.sg -> 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_goals_rule_aux: (simpset -> thm -> thm option) -> thm list -> thm -> thm wenzelm@15023: val rewrite_goal_rule: bool * bool * bool -> wenzelm@15023: (simpset -> thm -> thm option) -> simpset -> int -> thm -> thm wenzelm@15023: val asm_rewrite_goal_tac: bool * bool * bool -> wenzelm@15023: (simpset -> tactic) -> simpset -> int -> tactic wenzelm@15023: val simp_thm: bool * bool * bool -> simpset -> thm -> thm wenzelm@15023: val simp_cterm: bool * bool * bool -> simpset -> cterm -> thm berghofe@10413: end; berghofe@10413: wenzelm@15023: structure MetaSimplifier: META_SIMPLIFIER = berghofe@10413: struct berghofe@10413: wenzelm@15023: berghofe@10413: (** diagnostics **) berghofe@10413: berghofe@10413: exception SIMPLIFIER of string * thm; berghofe@10413: wenzelm@15023: val debug_simp = ref false; wenzelm@15023: val trace_simp = ref false; nipkow@11505: val simp_depth = ref 0; nipkow@13828: val simp_depth_limit = ref 1000; nipkow@11505: wenzelm@12603: local wenzelm@12603: wenzelm@12603: fun println a = wenzelm@15023: tracing (case ! simp_depth of 0 => a | n => enclose "[" "]" (string_of_int n) ^ a); nipkow@11505: nipkow@11505: fun prnt warn a = if warn then warning a else println a; wenzelm@15023: fun prtm warn a sg t = prnt warn (a ^ "\n" ^ Sign.string_of_term sg t); wenzelm@12603: fun prctm warn a t = prnt warn (a ^ "\n" ^ Display.string_of_cterm t); berghofe@10413: wenzelm@12603: in berghofe@10413: wenzelm@15023: fun debug warn a = if ! debug_simp then prnt warn a else (); wenzelm@15023: fun trace warn a = if ! trace_simp then prnt warn a else (); berghofe@10413: wenzelm@15023: fun debug_term warn a sign t = if ! debug_simp then prtm warn a sign t else (); wenzelm@15023: fun trace_term warn a sign t = if ! trace_simp then prtm warn a sign t else (); wenzelm@15023: fun trace_cterm warn a ct = if ! trace_simp then prctm warn a ct else (); wenzelm@15023: fun trace_thm a th = if ! trace_simp then prctm false a (Thm.cprop_of th) else (); nipkow@13569: berghofe@13607: fun trace_named_thm a (thm, name) = wenzelm@15023: if ! trace_simp then wenzelm@15023: prctm false (if name = "" then a else a ^ " " ^ quote name ^ ":") (Thm.cprop_of thm) wenzelm@15023: else (); wenzelm@15023: wenzelm@15023: fun warn_thm a = prctm true a o Thm.cprop_of; berghofe@10413: wenzelm@12603: end; berghofe@10413: berghofe@10413: berghofe@10413: wenzelm@15023: (** datatype simpset **) wenzelm@15023: wenzelm@15023: (* rewrite rules *) berghofe@10413: berghofe@13607: type rrule = {thm: thm, name: string, lhs: term, elhs: cterm, fo: bool, perm: bool}; wenzelm@15023: wenzelm@15023: (*thm: the rewrite rule; wenzelm@15023: name: name of theorem from which rewrite rule was extracted; wenzelm@15023: lhs: the left-hand side; wenzelm@15023: elhs: the etac-contracted lhs; wenzelm@15023: fo: use first-order matching; wenzelm@15023: perm: the rewrite rule is permutative; wenzelm@15023: 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@15023: in which case there is nothing better to do;*) berghofe@10413: berghofe@10413: fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = wenzelm@15023: Drule.eq_thm_prop (thm1, thm2); wenzelm@15023: wenzelm@15023: wenzelm@15023: (* congruences *) wenzelm@15023: wenzelm@15023: type cong = {thm: thm, lhs: cterm}; berghofe@10413: wenzelm@12603: fun eq_cong ({thm = thm1, ...}: cong, {thm = thm2, ...}: cong) = wenzelm@15023: Drule.eq_thm_prop (thm1, thm2); berghofe@10413: berghofe@10413: wenzelm@15023: (* solvers *) wenzelm@15023: wenzelm@15023: datatype solver = wenzelm@15023: Solver of wenzelm@15023: {name: string, wenzelm@15023: solver: thm list -> int -> tactic, wenzelm@15023: id: stamp}; wenzelm@15023: wenzelm@15023: fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()}; wenzelm@15023: wenzelm@15034: fun solver_name (Solver {name, ...}) = name; wenzelm@15023: fun solver ths (Solver {solver = tacf, ...}) = tacf ths; wenzelm@15023: fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); wenzelm@15023: val merge_solvers = gen_merge_lists eq_solver; wenzelm@15023: wenzelm@15023: wenzelm@15023: (* simplification sets and procedures *) 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@15023: depth: depth of conditional rewriting; 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@15023: mk_eq_True: thm -> thm option}; wenzelm@15023: wenzelm@15023: datatype simpset = wenzelm@15023: Simpset of wenzelm@15023: {rules: rrule Net.net, berghofe@10413: prems: thm list, berghofe@15249: bounds: int, wenzelm@15023: depth: int} * wenzelm@15023: {congs: (string * cong) 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, skalberg@15011: loop_tacs: (string * (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@15023: proc: Sign.sg -> simpset -> term -> thm option, wenzelm@15023: id: stamp}; wenzelm@15023: wenzelm@15023: fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = (id1 = id2); wenzelm@15023: wenzelm@15023: fun rep_ss (Simpset args) = args; berghofe@10413: wenzelm@15023: fun make_ss1 (rules, prems, bounds, depth) = wenzelm@15023: {rules = rules, prems = prems, bounds = bounds, depth = depth}; wenzelm@15023: wenzelm@15023: fun map_ss1 f {rules, prems, bounds, depth} = wenzelm@15023: make_ss1 (f (rules, prems, bounds, depth)); 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_simpset f (Simpset ({rules, prems, bounds, depth}, wenzelm@15023: {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers})) = wenzelm@15023: make_simpset (f ((rules, prems, bounds, depth), wenzelm@15023: (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers))); 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@15023: wenzelm@15023: (* print simpsets *) berghofe@10413: wenzelm@15023: fun print_ss ss = wenzelm@15023: let wenzelm@15034: val pretty_thms = map Display.pretty_thm; wenzelm@15023: wenzelm@15034: fun pretty_cong (name, th) = wenzelm@15034: Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, Display.pretty_thm th]; wenzelm@15023: fun pretty_proc (name, lhss) = wenzelm@15023: Pretty.big_list (name ^ ":") (map Display.pretty_cterm lhss); wenzelm@15034: wenzelm@15034: val Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...}) = ss; wenzelm@15034: val smps = map (#thm o #2) (Net.dest rules); wenzelm@15034: val cngs = map (fn (name, {thm, ...}) => (name, thm)) (#1 congs); wenzelm@15034: val prcs = Net.dest procs |> wenzelm@15034: map (fn (_, Proc {name, lhs, id, ...}) => ((name, lhs), id)) wenzelm@15034: |> partition_eq eq_snd wenzelm@15034: |> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps)) wenzelm@15034: |> Library.sort_wrt #1; wenzelm@15023: in wenzelm@15034: [Pretty.big_list "simplification rules:" (pretty_thms smps), wenzelm@15034: Pretty.big_list "simplification procedures:" (map pretty_proc prcs), wenzelm@15034: Pretty.big_list "congruences:" (map pretty_cong cngs), wenzelm@15088: Pretty.strs ("loopers:" :: map (quote o #1) loop_tacs), wenzelm@15088: Pretty.strs ("unsafe solvers:" :: map (quote o solver_name) (#1 solvers)), wenzelm@15088: Pretty.strs ("safe solvers:" :: map (quote o solver_name) (#2 solvers))] wenzelm@15023: |> Pretty.chunks |> Pretty.writeln nipkow@13828: end; berghofe@10413: wenzelm@15023: wenzelm@15023: (* empty simpsets *) wenzelm@15023: wenzelm@15023: local wenzelm@15023: wenzelm@15023: fun init_ss mk_rews termless subgoal_tac solvers = berghofe@15249: make_simpset ((Net.empty, [], 0, 0), wenzelm@15023: (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers)); wenzelm@15023: wenzelm@15023: val basic_mk_rews: mk_rews = wenzelm@15023: {mk = fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], wenzelm@15023: mk_cong = I, skalberg@15531: mk_sym = SOME o Drule.symmetric_fun, skalberg@15531: mk_eq_True = K NONE}; wenzelm@15023: wenzelm@15023: in wenzelm@15023: wenzelm@15023: val empty_ss = init_ss basic_mk_rews Term.termless (K (K no_tac)) ([], []); wenzelm@15023: wenzelm@15023: fun clear_ss (Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...})) = wenzelm@15023: init_ss mk_rews termless subgoal_tac solvers; wenzelm@15023: wenzelm@15023: end; wenzelm@15023: wenzelm@15023: wenzelm@15023: (* merge simpsets *) (*NOTE: ignores some fields of 2nd simpset*) skalberg@15011: wenzelm@15023: fun merge_ss (ss1, ss2) = wenzelm@15023: let wenzelm@15023: val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth}, wenzelm@15023: {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac, wenzelm@15023: loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; wenzelm@15023: val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = _}, wenzelm@15023: {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _, wenzelm@15023: loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; skalberg@15011: wenzelm@15023: val rules' = Net.merge (rules1, rules2, eq_rrule); wenzelm@15023: val prems' = gen_merge_lists Drule.eq_thm_prop prems1 prems2; berghofe@15249: val bounds' = Int.max (bounds1, bounds2); wenzelm@15023: val congs' = gen_merge_lists (eq_cong o pairself #2) congs1 congs2; wenzelm@15023: val weak' = merge_lists weak1 weak2; wenzelm@15023: val procs' = Net.merge (procs1, procs2, eq_proc); wenzelm@15023: val loop_tacs' = merge_alists loop_tacs1 loop_tacs2; wenzelm@15023: val unsafe_solvers' = merge_solvers unsafe_solvers1 unsafe_solvers2; wenzelm@15023: val solvers' = merge_solvers solvers1 solvers2; wenzelm@15023: in wenzelm@15023: make_simpset ((rules', prems', bounds', depth), ((congs', weak'), procs', wenzelm@15023: mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) wenzelm@15023: end; wenzelm@15023: wenzelm@15023: wenzelm@15023: (* simprocs *) wenzelm@15023: wenzelm@15023: exception SIMPROC_FAIL of string * exn; wenzelm@15023: wenzelm@15023: datatype simproc = Simproc of proc list; wenzelm@15023: wenzelm@15023: fun mk_simproc name lhss proc = wenzelm@15023: let val id = stamp () in wenzelm@15023: Simproc (lhss |> map (fn lhs => wenzelm@15023: Proc {name = name, lhs = lhs, proc = proc, id = id})) wenzelm@15023: end; wenzelm@15023: wenzelm@15023: fun simproc_i sg name = mk_simproc name o map (Thm.cterm_of sg o Logic.varify); wenzelm@15023: fun simproc sg name = simproc_i sg name o map (Sign.simple_read_term sg TypeInfer.logicT); wenzelm@15023: skalberg@15011: berghofe@10413: berghofe@10413: (** simpset operations **) berghofe@10413: wenzelm@15023: (* bounds and prems *) berghofe@10413: berghofe@15249: val incr_bounds = map_simpset1 (fn (rules, prems, bounds, depth) => berghofe@15249: (rules, prems, bounds + 1, depth)); berghofe@10413: wenzelm@15023: fun add_prems ths = map_simpset1 (fn (rules, prems, bounds, depth) => wenzelm@15023: (rules, ths @ prems, bounds, depth)); wenzelm@15023: wenzelm@15023: fun prems_of_ss (Simpset ({prems, ...}, _)) = prems; berghofe@10413: berghofe@10413: wenzelm@15023: (* addsimps *) berghofe@10413: wenzelm@15023: fun mk_rrule2 {thm, name, lhs, elhs, perm} = wenzelm@15023: let wenzelm@15023: val fo = Pattern.first_order (term_of elhs) orelse not (Pattern.pattern (term_of elhs)) wenzelm@15023: in {thm = thm, name = name, lhs = lhs, elhs = elhs, fo = fo, perm = perm} end; berghofe@10413: wenzelm@15023: fun insert_rrule quiet (ss, rrule as {thm, name, lhs, elhs, perm}) = wenzelm@15023: (trace_named_thm "Adding rewrite rule" (thm, name); wenzelm@15023: ss |> map_simpset1 (fn (rules, prems, bounds, depth) => wenzelm@15023: let wenzelm@15023: val rrule2 as {elhs, ...} = mk_rrule2 rrule; wenzelm@15023: val rules' = Net.insert_term ((term_of elhs, rrule2), rules, eq_rrule); wenzelm@15023: in (rules', prems, bounds, depth) end) wenzelm@15023: handle Net.INSERT => wenzelm@15023: (if quiet then () else warn_thm "Ignoring duplicate rewrite rule:" 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) = berghofe@10413: vperm (t, u) andalso eq_set (term_varnames t, term_varnames u); berghofe@10413: berghofe@10413: (* FIXME: it seems that the conditions on extra variables are too liberal if berghofe@10413: prems are nonempty: does solving the prems really guarantee instantiation of berghofe@10413: all its Vars? Better: a dynamic check each time a rule is applied. berghofe@10413: *) berghofe@10413: fun rewrite_rule_extra_vars prems elhs erhs = skalberg@15570: not (term_varnames erhs subset Library.foldl add_term_varnames (term_varnames elhs, prems)) berghofe@10413: orelse wenzelm@15023: not (term_tvars erhs subset (term_tvars elhs union List.concat (map term_tvars prems))); berghofe@10413: wenzelm@15023: (*simple test for looping rewrite rules and stupid orientations*) berghofe@10413: fun reorient sign 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 wenzelm@15023: exists (apl (lhs, Logic.occs)) (rhs :: prems) wenzelm@15023: orelse wenzelm@15023: null prems andalso Pattern.matches (Sign.tsig_of sign) (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@15023: val {sign, prop, ...} = Thm.rep_thm thm; wenzelm@15023: val prems = Logic.strip_imp_prems prop; wenzelm@15023: val concl = Drule.strip_imp_concl (Thm.cprop_of thm); wenzelm@15023: val (lhs, rhs) = Drule.dest_equals concl handle TERM _ => wenzelm@15023: raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm); wenzelm@15023: val (_, elhs) = Drule.dest_equals (Thm.cprop_of (Thm.eta_conversion lhs)); wenzelm@15023: val elhs = if elhs = lhs then lhs else elhs; (*share identical copies*) wenzelm@15023: val erhs = Pattern.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)); berghofe@10413: in (sign, 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@15023: let 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@15023: if term_varnames rhs subset term_varnames lhs andalso wenzelm@15023: term_tvars rhs subset term_tvars lhs then [rrule] wenzelm@15023: else mk_eq_True ss (thm2, name) @ [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@15023: let val (sign, 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 if reorient sign prems lhs rhs then wenzelm@15023: if reorient sign prems rhs lhs wenzelm@15023: then mk_eq_True ss (thm, name) wenzelm@15023: else wenzelm@15023: let val Simpset (_, {mk_rews = {mk_sym, ...}, ...}) = ss in wenzelm@15023: (case mk_sym thm of skalberg@15531: NONE => [] skalberg@15531: | SOME thm' => wenzelm@15023: let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm' wenzelm@15023: in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end) wenzelm@15023: 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) = skalberg@15570: List.concat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms); berghofe@10413: wenzelm@15023: fun orient_comb_simps comb mk_rrule (ss, thms) = wenzelm@15023: let wenzelm@15023: val rews = extract_rews (ss, thms); skalberg@15570: val rrules = List.concat (map mk_rrule rews); skalberg@15570: in Library.foldl comb (ss, rrules) end; berghofe@10413: wenzelm@15023: fun extract_safe_rrules (ss, thm) = skalberg@15570: List.concat (map (orient_rrule ss) (extract_rews (ss, [thm]))); berghofe@10413: wenzelm@15023: (*add rewrite rules explicitly; do not reorient!*) wenzelm@15023: fun ss addsimps thms = wenzelm@15023: orient_comb_simps (insert_rrule false) (mk_rrule ss) (ss, thms); berghofe@10413: berghofe@10413: wenzelm@15023: (* delsimps *) berghofe@10413: wenzelm@15023: fun del_rrule (ss, rrule as {thm, elhs, ...}) = wenzelm@15023: ss |> map_simpset1 (fn (rules, prems, bounds, depth) => wenzelm@15023: (Net.delete_term ((term_of elhs, rrule), rules, eq_rrule), prems, bounds, depth)) wenzelm@15023: handle Net.DELETE => (warn_thm "Rewrite rule not in simpset:" thm; ss); berghofe@10413: wenzelm@15023: fun ss delsimps thms = wenzelm@15023: orient_comb_simps del_rrule (map mk_rrule2 o mk_rrule ss) (ss, thms); wenzelm@15023: wenzelm@15023: 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 wenzelm@15023: null (findrep xs) andalso xs = ys andalso wenzelm@15023: (x, y) mem varpairs andalso wenzelm@15023: is_full_cong_prems prems (varpairs \ (x, y)) 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 wenzelm@15023: f = g andalso null (findrep (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@15023: val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm)) wenzelm@15023: handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm); wenzelm@15023: (*val lhs = Pattern.eta_contract lhs;*) skalberg@15570: val a = valOf (cong_name (head_of (term_of lhs))) handle Option => wenzelm@15023: raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm); wenzelm@15023: val (alist, weak) = congs; wenzelm@15023: val alist2 = overwrite_warn (alist, (a, {lhs = lhs, thm = thm})) wenzelm@15023: ("Overwriting congruence rule for " ^ quote a); wenzelm@15023: val weak2 = if is_full_cong thm then weak else a :: weak; wenzelm@15023: in ((alist2, weak2), 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@15023: (*val lhs = Pattern.eta_contract lhs;*) skalberg@15570: val a = valOf (cong_name (head_of lhs)) handle Option => wenzelm@15023: raise SIMPLIFIER ("Congruence must start with a constant", thm); wenzelm@15023: val (alist, _) = congs; skalberg@15570: val alist2 = List.filter (fn (x, _) => x <> a) alist; skalberg@15570: val weak2 = alist2 |> List.mapPartial (fn (a, {thm, ...}) => skalberg@15531: if is_full_cong thm then NONE else SOME a); wenzelm@15023: in ((alist2, weak2), 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@15023: local berghofe@10413: wenzelm@15023: fun add_proc (ss, proc as Proc {name, lhs, ...}) = wenzelm@15023: (trace_cterm false ("Adding simplification procedure " ^ quote name ^ " for") lhs; wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: (congs, Net.insert_term ((term_of lhs, proc), procs, eq_proc), 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@15023: fun del_proc (ss, proc as Proc {name, lhs, ...}) = wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15023: (congs, Net.delete_term ((term_of lhs, proc), procs, eq_proc), 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@15023: in berghofe@10413: skalberg@15570: val (op addsimprocs) = Library.foldl (fn (ss, Simproc procs) => Library.foldl add_proc (ss, procs)); skalberg@15570: val (op delsimprocs) = Library.foldl (fn (ss, Simproc procs) => Library.foldl del_proc (ss, procs)); berghofe@10413: wenzelm@15023: end; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_rews *) berghofe@10413: wenzelm@15023: local wenzelm@15023: nipkow@15199: fun map_mk_rews f = map_simpset2 (fn (congs, procs, {mk, mk_cong, mk_sym, mk_eq_True}, wenzelm@15023: termless, subgoal_tac, loop_tacs, solvers) => nipkow@15199: let val (mk', mk_cong', mk_sym', mk_eq_True') = f (mk, mk_cong, mk_sym, mk_eq_True) in nipkow@15199: (congs, procs, {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True'}, wenzelm@15023: termless, subgoal_tac, loop_tacs, solvers) wenzelm@15023: end); wenzelm@15023: wenzelm@15023: in berghofe@10413: nipkow@15199: fun ss setmksimps mk = ss |> map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True) => nipkow@15199: (mk, mk_cong, mk_sym, mk_eq_True)); wenzelm@15023: nipkow@15199: fun ss setmkcong mk_cong = ss |> map_mk_rews (fn (mk, _, mk_sym, mk_eq_True) => nipkow@15199: (mk, mk_cong, mk_sym, mk_eq_True)); berghofe@10413: nipkow@15199: fun ss setmksym mk_sym = ss |> map_mk_rews (fn (mk, mk_cong, _, mk_eq_True) => nipkow@15199: (mk, mk_cong, mk_sym, mk_eq_True)); berghofe@10413: nipkow@15199: fun ss setmkeqTrue mk_eq_True = ss |> map_mk_rews (fn (mk, mk_cong, mk_sym, _) => nipkow@15199: (mk, mk_cong, mk_sym, mk_eq_True)); 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@15023: 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@15023: 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, wenzelm@15023: overwrite_warn (loop_tacs, (name, tac)) ("Overwriting looper " ^ quote name), wenzelm@15023: solvers)); skalberg@15006: wenzelm@15023: fun ss delloop name = ss |> wenzelm@15023: map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => wenzelm@15034: let val loop_tacs' = filter_out (equal name o #1) loop_tacs in wenzelm@15034: if length loop_tacs <> length loop_tacs' then () wenzelm@15034: else warning ("No such looper in simpset: " ^ quote name); wenzelm@15034: (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs', solvers) wenzelm@15023: end); 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, wenzelm@15023: subgoal_tac, loop_tacs, (unsafe_solvers, merge_solvers solvers [solver]))); 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, wenzelm@15023: subgoal_tac, loop_tacs, (merge_solvers unsafe_solvers [solver], 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: 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@15023: val dest_eq = Drule.dest_equals o Thm.cprop_of; wenzelm@15023: val lhs_of = #1 o dest_eq; wenzelm@15023: val rhs_of = #2 o dest_eq; berghofe@10413: berghofe@10413: fun check_conv msg thm thm' = berghofe@10413: let berghofe@10413: val thm'' = transitive thm (transitive skalberg@15001: (symmetric (Drule.beta_eta_conversion (lhs_of thm'))) thm') skalberg@15531: in if msg then trace_thm "SUCCEEDED" thm' else (); SOME thm'' end berghofe@10413: handle THM _ => wenzelm@15023: let val {sign, prop = _ $ _ $ prop0, ...} = Thm.rep_thm thm in wenzelm@15023: trace_thm "Proved wrong thm (Check subgoaler?)" thm'; wenzelm@15023: trace_term false "Should have proved:" sign prop0; skalberg@15531: NONE berghofe@10413: end; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_procrule *) berghofe@10413: berghofe@10413: fun mk_procrule thm = wenzelm@15023: let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in wenzelm@15023: if rewrite_rule_extra_vars prems lhs rhs wenzelm@15023: then (warn_thm "Extra vars on rhs:" 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@15023: else if exists_Const (fn (c, _) => c mem weak) 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@15023: fun cond_skel (args as (congs, (lhs, rhs))) = wenzelm@15023: if term_varnames rhs subset term_varnames lhs then uncond_skel args berghofe@10413: else skel0; berghofe@10413: wenzelm@15023: fun incr_depth ss = wenzelm@15023: let wenzelm@15023: val ss' = ss |> map_simpset1 (fn (rules, prems, bounds, depth) => wenzelm@15023: (rules, prems, bounds, depth + 1)); wenzelm@15023: val Simpset ({depth = depth', ...}, _) = ss'; wenzelm@15023: in wenzelm@15023: if depth' > ! simp_depth_limit skalberg@15531: then (warning "simp_depth_limit exceeded - giving up"; NONE) wenzelm@15023: else wenzelm@15023: (if depth' mod 10 = 0 wenzelm@15023: then warning ("Simplification depth " ^ string_of_int depth') wenzelm@15023: else (); skalberg@15531: SOME ss') wenzelm@15023: end; wenzelm@15023: 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@15023: fun rewritec (prover, signt, maxt) ss t = berghofe@10413: let wenzelm@15023: val Simpset ({rules, ...}, {congs, procs, termless, ...}) = ss; berghofe@10413: val eta_thm = Thm.eta_conversion t; berghofe@10413: val eta_t' = rhs_of eta_thm; berghofe@10413: val eta_t = term_of eta_t'; berghofe@10413: val tsigt = Sign.tsig_of signt; berghofe@13607: fun rew {thm, name, lhs, elhs, fo, perm} = berghofe@10413: let berghofe@10413: val {sign, prop, maxidx, ...} = rep_thm thm; berghofe@10413: val (rthm, elhs') = if maxt = ~1 then (thm, elhs) berghofe@10413: else (Thm.incr_indexes (maxt+1) thm, Thm.cterm_incr_indexes (maxt+1) elhs); berghofe@10413: val insts = if fo then Thm.cterm_first_order_match (elhs', eta_t') berghofe@10413: else Thm.cterm_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'; berghofe@10413: val unconditional = (Logic.count_prems (prop',0) = 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')) berghofe@13607: then (trace_named_thm "Cannot apply permutative rewrite rule" (thm, name); skalberg@15531: trace_thm "Term does not become smaller:" thm'; NONE) berghofe@13607: else (trace_named_thm "Applying instance of rewrite rule" (thm, name); berghofe@10413: if unconditional berghofe@10413: then nipkow@13569: (trace_thm "Rewriting:" thm'; berghofe@10413: let val lr = Logic.dest_equals prop; skalberg@15531: val SOME thm'' = check_conv false eta_thm thm' skalberg@15531: in SOME (thm'', uncond_skel (congs, lr)) end) berghofe@10413: else nipkow@13569: (trace_thm "Trying to rewrite:" thm'; wenzelm@15023: case incr_depth ss of skalberg@15531: NONE => (trace_thm "FAILED - reached depth limit" thm'; NONE) skalberg@15531: | SOME ss' => wenzelm@15023: (case prover ss' thm' of skalberg@15531: NONE => (trace_thm "FAILED" thm'; NONE) skalberg@15531: | SOME thm2 => berghofe@10413: (case check_conv true 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 skalberg@15531: 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@15023: if Pattern.matches tsigt (Thm.term_of lhs, Thm.term_of t) then berghofe@10413: (debug_term false ("Trying procedure " ^ quote name ^ " on:") signt eta_t; wenzelm@13486: case transform_failure (curry SIMPROC_FAIL name) wenzelm@15023: (fn () => proc signt ss eta_t) () of skalberg@15531: NONE => (debug false "FAILED"; proc_rews ps) skalberg@15531: | SOME raw_thm => nipkow@13569: (trace_thm ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm; berghofe@10413: (case rews (mk_procrule raw_thm) of skalberg@15531: NONE => (trace_cterm true ("IGNORED result of simproc " ^ quote name ^ wenzelm@13486: " -- does not match") 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: berghofe@10413: fun congc (prover,signt,maxt) {thm=cong,lhs=lhs} t = wenzelm@14643: let val sign = Thm.sign_of_thm cong berghofe@10413: val rthm = if maxt = ~1 then cong else Thm.incr_indexes (maxt+1) cong; berghofe@10413: val rlhs = fst (Drule.dest_equals (Drule.strip_imp_concl (cprop_of rthm))); berghofe@10413: val insts = Thm.cterm_match (rlhs, t) berghofe@10413: (* Pattern.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); nipkow@13569: val unit = trace_thm "Applying congruence rule:" thm'; skalberg@15531: fun err (msg, thm) = (trace_thm msg thm; NONE) berghofe@10413: in case prover thm' of skalberg@15531: NONE => err ("Congruence proof failed. Could not prove", thm') skalberg@15531: | SOME thm2 => (case check_conv true (Drule.beta_eta_conversion t) thm2 of skalberg@15531: NONE => err ("Congruence proof failed. Should not have proved", thm2) skalberg@15531: | SOME thm2' => berghofe@12155: if op aconv (pairself term_of (dest_equals (cprop_of thm2'))) skalberg@15531: then NONE else SOME thm2') berghofe@10413: end; berghofe@10413: berghofe@10413: val (cA, (cB, cC)) = berghofe@10413: apsnd dest_equals (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@15023: fun bottomc ((simprem, useprem, mutsimp), prover, sign, 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@15023: (case rewritec (prover, sign, maxidx) ss (rhs_of thm1) of skalberg@15531: SOME (thm2, skel2) => berghofe@13607: transitive2 (transitive thm1 thm2) wenzelm@15023: (botc skel2 ss (rhs_of thm2)) skalberg@15531: | NONE => some) skalberg@15531: | NONE => wenzelm@15023: (case rewritec (prover, sign, maxidx) ss t of skalberg@15531: SOME (thm2, skel2) => transitive2 thm2 wenzelm@15023: (botc skel2 ss (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 berghofe@10413: Abs (a, T, t) => wenzelm@15023: let skalberg@15531: val (v, t') = Thm.dest_abs (SOME ("." ^ a ^ "." ^ string_of_int bounds)) t0; berghofe@15249: val ss' = incr_bounds 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@15023: in case subc skel0 ss (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 => berghofe@10413: (case assoc_string (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@15023: val thm = congc (prover ss, sign, maxidx) cong t0; skalberg@15570: val t = getOpt (Option.map rhs_of thm, t0); 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)) berghofe@10413: end handle TERM _ => error "congc result" berghofe@10413: | 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 skalberg@15531: "Cannot add premise as rewrite rule because it contains (type) unknowns:" 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 = skalberg@15570: Library.foldl (insert_rrule true) (ss, List.concat rrss) |> add_prems (List.mapPartial I asms) berghofe@10413: berghofe@13607: and disch r (prem, eq) = berghofe@13607: let berghofe@13607: val (lhs, rhs) = dest_eq 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 berghofe@13607: val (prem', concl) = dest_implies lhs; berghofe@13607: val (prem'', _) = 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@15023: | rebuild (prem :: prems) concl (rrs :: rrss) (asm :: asms) ss eq = berghofe@13607: let wenzelm@15023: val ss' = add_rrules (rev rrss, rev asms) ss; berghofe@13607: val concl' = skalberg@15570: Drule.mk_implies (prem, getOpt (Option.map rhs_of eq, concl)); skalberg@15570: val dprem = Option.map (curry (disch false) prem) wenzelm@15023: in case rewritec (prover, sign, maxidx) ss' concl' of skalberg@15531: NONE => rebuild prems concl' rrss asms ss (dprem eq) skalberg@15570: | SOME (eq', _) => transitive2 (Library.foldl (disch false o swap) skalberg@15570: (valOf (transitive3 (dprem eq) eq'), prems)) wenzelm@15023: (mut_impc0 (rev prems) (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 = skalberg@15570: transitive1 (Library.foldl (fn (eq2, (eq1, prem)) => transitive1 eq1 skalberg@15570: (Option.map (curry (disch false) prem) eq2)) (NONE, eqns ~~ prems')) 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 berghofe@13607: val prem' = rhs_of eqn; berghofe@13607: val tprems = map term_of prems; skalberg@15570: val i = 1 + Library.foldl Int.max (~1, map (fn p => berghofe@13607: find_index_eq p tprems) (#hyps (rep_thm eqn))); wenzelm@15023: val (rrs', asm') = rules_of_prem ss prem' berghofe@13607: in mut_impc prems concl rrss asms (prem' :: prems') skalberg@15570: (rrs' :: rrss') (asm' :: asms') (SOME (Library.foldr (disch true) skalberg@15570: (Library.take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies skalberg@15570: (Library.drop (i, prems), concl))))) :: eqns) ss (length prems') ~1 berghofe@13607: end berghofe@13607: wenzelm@15023: (*legacy code - only for backwards compatibility*) wenzelm@15023: and nonmut_impc ct ss = berghofe@13607: let val (prem, conc) = dest_implies ct; skalberg@15531: val thm1 = if simprem then botc skel0 ss prem else NONE; skalberg@15570: val prem1 = getOpt (Option.map rhs_of thm1, prem); 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 skalberg@15531: | SOME thm1' => SOME (Drule.imp_cong' thm1' (reflexive conc))) skalberg@15531: | SOME thm2 => berghofe@13607: let val thm2' = disch false (prem1, thm2) berghofe@10413: in (case thm1 of skalberg@15531: NONE => SOME thm2' skalberg@15531: | SOME thm1' => skalberg@15531: SOME (transitive (Drule.imp_cong' 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@15023: fun rewrite_cterm mode prover ss ct = wenzelm@15023: let wenzelm@15023: val Simpset ({depth, ...}, _) = ss; wenzelm@15023: val {sign, t, maxidx, ...} = Thm.rep_cterm ct; wenzelm@15023: in wenzelm@15023: trace_cterm false "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" ct; wenzelm@15023: simp_depth := depth; wenzelm@15023: bottomc (mode, prover, sign, maxidx) ss ct wenzelm@15023: end handle THM (s, _, thms) => berghofe@10413: error ("Exception THM was raised in simplifier:\n" ^ s ^ "\n" ^ wenzelm@11886: Pretty.string_of (Display.pretty_thms thms)); berghofe@10413: wenzelm@11760: (*Rewrite a cterm*) wenzelm@11767: fun rewrite_aux _ _ [] = (fn ct => Thm.reflexive ct) wenzelm@15023: | rewrite_aux prover full thms = wenzelm@15023: rewrite_cterm (full, false, false) prover (empty_ss addsimps thms); wenzelm@11672: berghofe@10413: (*Rewrite a theorem*) wenzelm@11767: fun simplify_aux _ _ [] = (fn th => th) wenzelm@11767: | simplify_aux prover full thms = wenzelm@15023: Drule.fconv_rule (rewrite_cterm (full, false, false) prover (empty_ss addsimps thms)); berghofe@10413: wenzelm@15023: (*simple term rewriting -- no proof*) wenzelm@15023: fun rewrite_term sg rules procs = wenzelm@15023: Pattern.rewrite_term (Sign.tsig_of sg) (map decomp_simp' rules) procs; wenzelm@15023: wenzelm@15023: fun rewrite_thm mode prover ss = Drule.fconv_rule (rewrite_cterm mode prover ss); berghofe@10413: berghofe@10413: (*Rewrite the subgoals of a proof state (represented by a theorem) *) berghofe@10413: fun rewrite_goals_rule_aux _ [] th = th berghofe@10413: | rewrite_goals_rule_aux prover thms th = skalberg@15001: Drule.fconv_rule (Drule.goals_conv (K true) (rewrite_cterm (true, true, false) prover wenzelm@15023: (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 = berghofe@10413: if 0 < i andalso i <= nprems_of thm skalberg@15011: then Drule.fconv_rule (Drule.goals_conv (fn j => j=i) (rewrite_cterm mode prover ss)) thm berghofe@10413: else raise THM("rewrite_goal_rule",i,[thm]); berghofe@10413: wenzelm@15023: (*Rewrite subgoal i only. SELECT_GOAL avoids inefficiencies in goals_conv.*) wenzelm@15023: fun asm_rewrite_goal_tac mode prover_tac ss = wenzelm@15023: SELECT_GOAL wenzelm@15023: (PRIMITIVE (rewrite_goal_rule mode (SINGLE o prover_tac) ss 1)); wenzelm@12783: wenzelm@15023: skalberg@15006: wenzelm@15023: (** simplification tactics and rules **) skalberg@15006: wenzelm@15023: fun solve_all_tac solvers ss = skalberg@15006: let wenzelm@15023: val Simpset (_, {subgoal_tac, ...}) = ss; wenzelm@15023: val solve_tac = subgoal_tac (set_solvers solvers ss) THEN_ALL_NEW (K no_tac); wenzelm@15023: in DEPTH_SOLVE (solve_tac 1) end; skalberg@15006: wenzelm@15023: (*NOTE: may instantiate unknowns that appear also in other subgoals*) wenzelm@15023: fun generic_simp_tac safe mode ss = wenzelm@15023: let wenzelm@15023: val Simpset ({prems, ...}, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = ss; wenzelm@15023: val loop_tac = FIRST' (map #2 loop_tacs); wenzelm@15023: val solve_tac = FIRST' (map (solver prems) (if safe then solvers else unsafe_solvers)); skalberg@15006: wenzelm@15023: fun simp_loop_tac i = wenzelm@15023: asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN wenzelm@15023: (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); wenzelm@15023: in simp_loop_tac end; skalberg@15006: wenzelm@15023: fun simp rew mode ss thm = skalberg@15006: let wenzelm@15023: val Simpset (_, {solvers = (unsafe_solvers, _), ...}) = ss; wenzelm@15023: val tacf = solve_all_tac unsafe_solvers; skalberg@15570: fun prover s th = Option.map #1 (Seq.pull (tacf s th)); skalberg@15011: in rew mode prover ss thm end; skalberg@15006: skalberg@15006: val simp_thm = simp rewrite_thm; skalberg@15006: val simp_cterm = simp rewrite_cterm; skalberg@15006: berghofe@10413: end; berghofe@10413: wenzelm@11672: structure BasicMetaSimplifier: BASIC_META_SIMPLIFIER = MetaSimplifier; wenzelm@11672: open BasicMetaSimplifier;