berghofe@10413: (* Title: Pure/meta_simplifier.ML berghofe@10413: ID: $Id$ wenzelm@11672: Author: Tobias Nipkow and Stefan Berghofer wenzelm@12783: License: GPL (GNU GENERAL PUBLIC LICENSE) berghofe@10413: wenzelm@11672: Meta-level Simplification. berghofe@10413: *) berghofe@10413: wenzelm@11672: signature BASIC_META_SIMPLIFIER = wenzelm@11672: sig wenzelm@11672: val trace_simp: bool ref wenzelm@11672: val debug_simp: bool ref nipkow@13828: val simp_depth_limit: int ref wenzelm@11672: end; wenzelm@11672: berghofe@10413: signature META_SIMPLIFIER = berghofe@10413: sig wenzelm@11672: include BASIC_META_SIMPLIFIER berghofe@10413: exception SIMPLIFIER of string * thm wenzelm@13486: exception SIMPROC_FAIL of string * exn berghofe@10413: type meta_simpset wenzelm@12603: val dest_mss : meta_simpset -> berghofe@10413: {simps: thm list, congs: thm list, procs: (string * cterm list) list} berghofe@10413: val empty_mss : meta_simpset wenzelm@12603: val clear_mss : meta_simpset -> meta_simpset wenzelm@12603: val merge_mss : meta_simpset * meta_simpset -> meta_simpset berghofe@10413: val add_simps : meta_simpset * thm list -> meta_simpset berghofe@10413: val del_simps : meta_simpset * thm list -> meta_simpset berghofe@10413: val mss_of : thm list -> meta_simpset berghofe@10413: val add_congs : meta_simpset * thm list -> meta_simpset berghofe@10413: val del_congs : meta_simpset * thm list -> meta_simpset wenzelm@12603: val add_simprocs : meta_simpset * berghofe@10413: (string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list berghofe@10413: -> meta_simpset wenzelm@12603: val del_simprocs : meta_simpset * berghofe@10413: (string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list berghofe@10413: -> meta_simpset berghofe@10413: val add_prems : meta_simpset * thm list -> meta_simpset berghofe@10413: val prems_of_mss : meta_simpset -> thm list berghofe@10413: val set_mk_rews : meta_simpset * (thm -> thm list) -> meta_simpset berghofe@10413: val set_mk_sym : meta_simpset * (thm -> thm option) -> meta_simpset berghofe@10413: val set_mk_eq_True : meta_simpset * (thm -> thm option) -> meta_simpset skalberg@14242: val get_mk_rews : meta_simpset -> thm -> thm list skalberg@14242: val get_mk_sym : meta_simpset -> thm -> thm option skalberg@14242: val get_mk_eq_True : meta_simpset -> thm -> thm option berghofe@10413: val set_termless : meta_simpset * (term * term -> bool) -> meta_simpset wenzelm@12779: val beta_eta_conversion: cterm -> thm wenzelm@11672: val rewrite_cterm: bool * bool * bool -> wenzelm@11672: (meta_simpset -> thm -> thm option) -> meta_simpset -> cterm -> thm berghofe@11736: val goals_conv : (int -> bool) -> (cterm -> thm) -> cterm -> thm berghofe@11736: val forall_conv : (cterm -> thm) -> cterm -> thm berghofe@11736: val fconv_rule : (cterm -> thm) -> thm -> thm wenzelm@11767: val rewrite_aux : (meta_simpset -> thm -> thm option) -> bool -> thm list -> cterm -> thm wenzelm@11767: val simplify_aux : (meta_simpset -> thm -> thm option) -> bool -> thm list -> thm -> thm berghofe@10413: val rewrite_thm : bool * bool * bool berghofe@10413: -> (meta_simpset -> thm -> thm option) berghofe@10413: -> meta_simpset -> thm -> thm berghofe@10413: val rewrite_goals_rule_aux: (meta_simpset -> thm -> thm option) -> thm list -> thm -> thm berghofe@10413: val rewrite_goal_rule : bool* bool * bool berghofe@10413: -> (meta_simpset -> thm -> thm option) berghofe@10413: -> meta_simpset -> int -> thm -> thm berghofe@13196: val rewrite_term: Sign.sg -> thm list -> (term -> term option) list -> term -> term berghofe@10413: end; berghofe@10413: berghofe@10413: structure MetaSimplifier : META_SIMPLIFIER = berghofe@10413: struct berghofe@10413: berghofe@10413: (** diagnostics **) berghofe@10413: berghofe@10413: exception SIMPLIFIER of string * thm; wenzelm@13486: exception SIMPROC_FAIL of string * exn; berghofe@10413: 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@12603: tracing ((case ! simp_depth of 0 => "" | n => "[" ^ string_of_int n ^ "]") ^ a); nipkow@11505: nipkow@11505: fun prnt warn a = if warn then warning a else println a; wenzelm@12603: fun prtm warn a sign t = prnt warn (a ^ "\n" ^ Sign.string_of_term sign 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@12603: fun prthm warn a = prctm warn a o Thm.cprop_of; berghofe@10413: berghofe@10413: val trace_simp = ref false; berghofe@10413: val debug_simp = ref false; berghofe@10413: berghofe@10413: fun trace warn a = if !trace_simp then prnt warn a else (); berghofe@10413: fun debug warn a = if !debug_simp then prnt warn a else (); berghofe@10413: berghofe@10413: fun trace_term warn a sign t = if !trace_simp then prtm warn a sign t else (); berghofe@10413: fun trace_cterm warn a t = if !trace_simp then prctm warn a t else (); berghofe@10413: fun debug_term warn a sign t = if !debug_simp then prtm warn a sign t else (); berghofe@10413: nipkow@13569: fun trace_thm a thm = berghofe@10413: let val {sign, prop, ...} = rep_thm thm nipkow@13569: in trace_term false a sign prop end; nipkow@13569: berghofe@13607: fun trace_named_thm a (thm, name) = berghofe@13607: trace_thm (a ^ (if name = "" then "" else " " ^ quote name) ^ ":") thm; berghofe@10413: wenzelm@12603: end; berghofe@10413: berghofe@10413: berghofe@10413: (** meta simp sets **) berghofe@10413: berghofe@10413: (* basic components *) berghofe@10413: berghofe@13607: type rrule = {thm: thm, name: string, lhs: term, elhs: cterm, fo: bool, perm: bool}; berghofe@10413: (* thm: the rewrite rule berghofe@13607: name: name of theorem from which rewrite rule was extracted berghofe@10413: lhs: the left-hand side berghofe@10413: elhs: the etac-contracted lhs. berghofe@10413: fo: use first-order matching berghofe@10413: perm: the rewrite rule is permutative wenzelm@12603: Remarks: berghofe@10413: - elhs is used for matching, berghofe@10413: 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), berghofe@10413: in which case fo-matching is complete, berghofe@10413: or elhs is not a pattern, berghofe@10413: in which case there is nothing better to do. berghofe@10413: *) berghofe@10413: type cong = {thm: thm, lhs: cterm}; berghofe@10413: type simproc = berghofe@10413: {name: string, proc: Sign.sg -> thm list -> term -> thm option, lhs: cterm, id: stamp}; berghofe@10413: berghofe@10413: fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = berghofe@10413: #prop (rep_thm thm1) aconv #prop (rep_thm thm2); berghofe@10413: wenzelm@12603: fun eq_cong ({thm = thm1, ...}: cong, {thm = thm2, ...}: cong) = berghofe@10413: #prop (rep_thm thm1) aconv #prop (rep_thm thm2); berghofe@10413: berghofe@10413: fun eq_prem (thm1, thm2) = berghofe@10413: #prop (rep_thm thm1) aconv #prop (rep_thm thm2); berghofe@10413: berghofe@10413: fun eq_simproc ({id = s1, ...}:simproc, {id = s2, ...}:simproc) = (s1 = s2); berghofe@10413: berghofe@10413: fun mk_simproc (name, proc, lhs, id) = berghofe@10413: {name = name, proc = proc, lhs = lhs, id = id}; berghofe@10413: berghofe@10413: berghofe@10413: (* datatype mss *) berghofe@10413: berghofe@10413: (* berghofe@10413: A "mss" contains data needed during conversion: berghofe@10413: rules: discrimination net of rewrite rules; 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); berghofe@10413: bounds: names of bound variables already used berghofe@10413: (for generating new names when rewriting under lambda abstractions); berghofe@10413: prems: current premises; berghofe@10413: mk_rews: mk: turns simplification thms into rewrite rules; berghofe@10413: mk_sym: turns == around; (needs Drule!) berghofe@10413: mk_eq_True: turns P into P == True - logic specific; berghofe@10413: termless: relation for ordered rewriting; nipkow@11504: depth: depth of conditional rewriting; berghofe@10413: *) berghofe@10413: berghofe@10413: datatype meta_simpset = berghofe@10413: Mss of { berghofe@10413: rules: rrule Net.net, berghofe@10413: congs: (string * cong) list * string list, berghofe@10413: procs: simproc Net.net, berghofe@10413: bounds: string list, berghofe@10413: prems: thm list, berghofe@10413: mk_rews: {mk: thm -> thm list, berghofe@10413: mk_sym: thm -> thm option, berghofe@10413: mk_eq_True: thm -> thm option}, nipkow@11504: termless: term * term -> bool, nipkow@11504: depth: int}; berghofe@10413: nipkow@11504: fun mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless, depth) = berghofe@10413: Mss {rules = rules, congs = congs, procs = procs, bounds = bounds, nipkow@11504: prems=prems, mk_rews=mk_rews, termless=termless, depth=depth}; berghofe@10413: nipkow@11504: fun upd_rules(Mss{rules,congs,procs,bounds,prems,mk_rews,termless,depth}, rules') = nipkow@11504: mk_mss(rules',congs,procs,bounds,prems,mk_rews,termless,depth); berghofe@10413: berghofe@10413: val empty_mss = berghofe@10413: let val mk_rews = {mk = K [], mk_sym = K None, mk_eq_True = K None} nipkow@11504: in mk_mss (Net.empty, ([], []), Net.empty, [], [], mk_rews, Term.termless, 0) end; berghofe@10413: berghofe@10413: fun clear_mss (Mss {mk_rews, termless, ...}) = nipkow@11504: mk_mss (Net.empty, ([], []), Net.empty, [], [], mk_rews, termless,0); berghofe@10413: nipkow@11504: fun incr_depth(Mss{rules,congs,procs,bounds,prems,mk_rews,termless,depth}) = nipkow@13828: let val depth1 = depth+1 nipkow@13828: in if depth1 > !simp_depth_limit nipkow@13828: then (warning "simp_depth_limit exceeded - giving up"; None) kleing@14040: else (if depth1 mod 10 = 0 nipkow@13828: then warning("Simplification depth " ^ string_of_int depth1) nipkow@13828: else (); nipkow@13828: Some(mk_mss(rules,congs,procs,bounds,prems,mk_rews,termless,depth1)) nipkow@13828: ) nipkow@13828: end; berghofe@10413: berghofe@10413: berghofe@10413: (** simpset operations **) berghofe@10413: berghofe@10413: (* term variables *) berghofe@10413: berghofe@10413: val add_term_varnames = foldl_aterms (fn (xs, Var (x, _)) => ins_ix (x, xs) | (xs, _) => xs); berghofe@10413: fun term_varnames t = add_term_varnames ([], t); berghofe@10413: berghofe@10413: berghofe@10413: (* dest_mss *) berghofe@10413: berghofe@10413: fun dest_mss (Mss {rules, congs, procs, ...}) = berghofe@10413: {simps = map (fn (_, {thm, ...}) => thm) (Net.dest rules), berghofe@10413: congs = map (fn (_, {thm, ...}) => thm) (fst congs), berghofe@10413: procs = berghofe@10413: map (fn (_, {name, lhs, id, ...}) => ((name, lhs), id)) (Net.dest procs) berghofe@10413: |> partition_eq eq_snd berghofe@10413: |> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps)) berghofe@10413: |> Library.sort_wrt #1}; berghofe@10413: berghofe@10413: wenzelm@12603: (* merge_mss *) (*NOTE: ignores mk_rews, termless and depth of 2nd mss*) berghofe@10413: berghofe@10413: fun merge_mss berghofe@10413: (Mss {rules = rules1, congs = (congs1,weak1), procs = procs1, nipkow@11504: bounds = bounds1, prems = prems1, mk_rews, termless, depth}, berghofe@10413: Mss {rules = rules2, congs = (congs2,weak2), procs = procs2, berghofe@10413: bounds = bounds2, prems = prems2, ...}) = berghofe@10413: mk_mss berghofe@10413: (Net.merge (rules1, rules2, eq_rrule), wenzelm@12285: (gen_merge_lists (eq_cong o pairself snd) congs1 congs2, berghofe@10413: merge_lists weak1 weak2), berghofe@10413: Net.merge (procs1, procs2, eq_simproc), berghofe@10413: merge_lists bounds1 bounds2, wenzelm@12285: gen_merge_lists eq_prem prems1 prems2, nipkow@11504: mk_rews, termless, depth); berghofe@10413: berghofe@10413: berghofe@10413: (* add_simps *) berghofe@10413: berghofe@13607: fun mk_rrule2{thm, name, lhs, elhs, perm} = berghofe@10413: let val fo = Pattern.first_order (term_of elhs) orelse not(Pattern.pattern (term_of elhs)) berghofe@13607: in {thm=thm, name=name, lhs=lhs, elhs=elhs, fo=fo, perm=perm} end berghofe@10413: berghofe@13607: fun insert_rrule quiet (mss as Mss {rules,...}, berghofe@13607: rrule as {thm,name,lhs,elhs,perm}) = berghofe@13607: (trace_named_thm "Adding rewrite rule" (thm, name); berghofe@10413: let val rrule2 as {elhs,...} = mk_rrule2 rrule berghofe@10413: val rules' = Net.insert_term ((term_of elhs, rrule2), rules, eq_rrule) berghofe@10413: in upd_rules(mss,rules') end berghofe@13607: handle Net.INSERT => if quiet then mss else berghofe@10413: (prthm true "Ignoring duplicate rewrite rule:" thm; mss)); 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 = berghofe@10413: not (term_varnames erhs subset foldl add_term_varnames (term_varnames elhs, prems)) berghofe@10413: orelse berghofe@10413: not ((term_tvars erhs) subset berghofe@10413: (term_tvars elhs union List.concat(map term_tvars prems))); berghofe@10413: berghofe@10413: (*Simple test for looping rewrite rules and stupid orientations*) berghofe@10413: fun reorient sign prems lhs rhs = berghofe@10413: rewrite_rule_extra_vars prems lhs rhs berghofe@10413: orelse berghofe@10413: is_Var (head_of lhs) berghofe@10413: orelse berghofe@10413: (exists (apl (lhs, Logic.occs)) (rhs :: prems)) berghofe@10413: orelse berghofe@10413: (null prems andalso berghofe@10413: Pattern.matches (#tsig (Sign.rep_sg 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 berghofe@10413: the rhs is an instance of the lhs. Example: ?m < ?n ==> f(?n) == f(?m)*) berghofe@10413: orelse berghofe@10413: (is_Const lhs andalso not(is_Const rhs)) berghofe@10413: berghofe@10413: fun decomp_simp thm = berghofe@10413: let val {sign, prop, ...} = rep_thm thm; berghofe@10413: val prems = Logic.strip_imp_prems prop; berghofe@10413: val concl = Drule.strip_imp_concl (cprop_of thm); berghofe@10413: val (lhs, rhs) = Drule.dest_equals concl handle TERM _ => berghofe@10413: raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm) berghofe@10413: val elhs = snd (Drule.dest_equals (cprop_of (Thm.eta_conversion lhs))); berghofe@10413: val elhs = if elhs=lhs then lhs else elhs (* try to share *) berghofe@10413: val erhs = Pattern.eta_contract (term_of rhs); berghofe@10413: val perm = var_perm (term_of elhs, erhs) andalso not (term_of elhs aconv erhs) berghofe@10413: andalso 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: berghofe@13607: fun mk_eq_True (Mss{mk_rews={mk_eq_True,...},...}) (thm, name) = berghofe@10413: case mk_eq_True thm of berghofe@10413: None => [] berghofe@13607: | Some eq_True => berghofe@13607: let val (_,_,lhs,elhs,_,_) = decomp_simp eq_True berghofe@13607: in [{thm=eq_True, name=name, lhs=lhs, elhs=elhs, perm=false}] end; berghofe@10413: berghofe@10413: (* create the rewrite rule and possibly also the ==True variant, berghofe@10413: in case there are extra vars on the rhs *) berghofe@13607: fun rrule_eq_True(thm,name,lhs,elhs,rhs,mss,thm2) = berghofe@13607: let val rrule = {thm=thm, name=name, lhs=lhs, elhs=elhs, perm=false} berghofe@10413: in if (term_varnames rhs) subset (term_varnames lhs) andalso berghofe@10413: (term_tvars rhs) subset (term_tvars lhs) berghofe@10413: then [rrule] berghofe@13607: else mk_eq_True mss (thm2, name) @ [rrule] berghofe@10413: end; berghofe@10413: berghofe@13607: fun mk_rrule mss (thm, name) = berghofe@10413: let val (_,prems,lhs,elhs,rhs,perm) = decomp_simp thm berghofe@13607: in if perm then [{thm=thm, name=name, lhs=lhs, elhs=elhs, perm=true}] else berghofe@10413: (* weak test for loops: *) berghofe@10413: if rewrite_rule_extra_vars prems lhs rhs orelse berghofe@10413: is_Var (term_of elhs) berghofe@13607: then mk_eq_True mss (thm, name) berghofe@13607: else rrule_eq_True(thm,name,lhs,elhs,rhs,mss,thm) berghofe@10413: end; berghofe@10413: berghofe@13607: fun orient_rrule mss (thm, name) = berghofe@10413: let val (sign,prems,lhs,elhs,rhs,perm) = decomp_simp thm berghofe@13607: in if perm then [{thm=thm, name=name, lhs=lhs, elhs=elhs, perm=true}] berghofe@10413: else if reorient sign prems lhs rhs berghofe@10413: then if reorient sign prems rhs lhs berghofe@13607: then mk_eq_True mss (thm, name) berghofe@10413: else let val Mss{mk_rews={mk_sym,...},...} = mss berghofe@10413: in case mk_sym thm of berghofe@10413: None => [] berghofe@10413: | Some thm' => berghofe@10413: let val (_,_,lhs',elhs',rhs',_) = decomp_simp thm' berghofe@13607: in rrule_eq_True(thm',name,lhs',elhs',rhs',mss,thm) end berghofe@10413: end berghofe@13607: else rrule_eq_True(thm,name,lhs,elhs,rhs,mss,thm) berghofe@10413: end; berghofe@10413: berghofe@13607: fun extract_rews(Mss{mk_rews = {mk,...},...},thms) = berghofe@13607: flat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms); berghofe@10413: berghofe@10413: fun orient_comb_simps comb mk_rrule (mss,thms) = berghofe@10413: let val rews = extract_rews(mss,thms) berghofe@10413: val rrules = flat (map mk_rrule rews) berghofe@10413: in foldl comb (mss,rrules) end berghofe@10413: berghofe@10413: (* Add rewrite rules explicitly; do not reorient! *) berghofe@10413: fun add_simps(mss,thms) = berghofe@13607: orient_comb_simps (insert_rrule false) (mk_rrule mss) (mss,thms); berghofe@10413: berghofe@13607: fun mss_of thms = foldl (insert_rrule false) (empty_mss, flat berghofe@13607: (map (fn thm => mk_rrule empty_mss (thm, Thm.name_of_thm thm)) thms)); berghofe@10413: berghofe@10413: fun extract_safe_rrules(mss,thm) = berghofe@10413: flat (map (orient_rrule mss) (extract_rews(mss,[thm]))); berghofe@10413: berghofe@10413: (* del_simps *) berghofe@10413: berghofe@10413: fun del_rrule(mss as Mss {rules,...}, berghofe@10413: rrule as {thm, elhs, ...}) = berghofe@10413: (upd_rules(mss, Net.delete_term ((term_of elhs, rrule), rules, eq_rrule)) berghofe@10413: handle Net.DELETE => berghofe@10413: (prthm true "Rewrite rule not in simpset:" thm; mss)); berghofe@10413: berghofe@10413: fun del_simps(mss,thms) = berghofe@10413: orient_comb_simps del_rrule (map mk_rrule2 o mk_rrule mss) (mss,thms); berghofe@10413: berghofe@10413: berghofe@10413: (* add_congs *) berghofe@10413: berghofe@10413: fun is_full_cong_prems [] varpairs = null varpairs berghofe@10413: | is_full_cong_prems (p::prems) varpairs = berghofe@10413: (case Logic.strip_assums_concl p of berghofe@10413: Const("==",_) $ lhs $ rhs => berghofe@10413: let val (x,xs) = strip_comb lhs and (y,ys) = strip_comb rhs berghofe@10413: in is_Var x andalso forall is_Bound xs andalso berghofe@10413: null(findrep(xs)) andalso xs=ys andalso berghofe@10413: (x,y) mem varpairs andalso berghofe@10413: is_full_cong_prems prems (varpairs\(x,y)) berghofe@10413: end berghofe@10413: | _ => false); berghofe@10413: berghofe@10413: fun is_full_cong thm = berghofe@10413: let val prems = prems_of thm berghofe@10413: and concl = concl_of thm berghofe@10413: val (lhs,rhs) = Logic.dest_equals concl berghofe@10413: val (f,xs) = strip_comb lhs berghofe@10413: and (g,ys) = strip_comb rhs berghofe@10413: in berghofe@10413: f=g andalso null(findrep(xs@ys)) andalso length xs = length ys andalso berghofe@10413: is_full_cong_prems prems (xs ~~ ys) berghofe@10413: end berghofe@10413: ballarin@13835: fun cong_name (Const (a, _)) = Some a ballarin@13835: | cong_name (Free (a, _)) = Some ("Free: " ^ a) ballarin@13835: | cong_name _ = None; ballarin@13835: nipkow@11504: fun add_cong (Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, thm) = berghofe@10413: let berghofe@10413: val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (cprop_of thm)) handle TERM _ => berghofe@10413: raise SIMPLIFIER ("Congruence not a meta-equality", thm); berghofe@10413: (* val lhs = Pattern.eta_contract lhs; *) ballarin@13835: val a = (case cong_name (head_of (term_of lhs)) of ballarin@13835: Some a => a ballarin@13835: | None => ballarin@13835: raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm)); berghofe@10413: val (alist,weak) = congs berghofe@10413: val alist2 = overwrite_warn (alist, (a,{lhs=lhs, thm=thm})) berghofe@10413: ("Overwriting congruence rule for " ^ quote a); berghofe@10413: val weak2 = if is_full_cong thm then weak else a::weak berghofe@10413: in nipkow@11504: mk_mss (rules,(alist2,weak2),procs,bounds,prems,mk_rews,termless,depth) berghofe@10413: end; berghofe@10413: berghofe@10413: val (op add_congs) = foldl add_cong; berghofe@10413: berghofe@10413: berghofe@10413: (* del_congs *) berghofe@10413: nipkow@11504: fun del_cong (Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, thm) = berghofe@10413: let berghofe@10413: val (lhs, _) = Logic.dest_equals (concl_of thm) handle TERM _ => berghofe@10413: raise SIMPLIFIER ("Congruence not a meta-equality", thm); berghofe@10413: (* val lhs = Pattern.eta_contract lhs; *) ballarin@13835: val a = (case cong_name (head_of lhs) of ballarin@13835: Some a => a ballarin@13835: | None => ballarin@13835: raise SIMPLIFIER ("Congruence must start with a constant", thm)); berghofe@10413: val (alist,_) = congs berghofe@10413: val alist2 = filter (fn (x,_)=> x<>a) alist berghofe@10413: val weak2 = mapfilter (fn(a,{thm,...}) => if is_full_cong thm then None berghofe@10413: else Some a) berghofe@10413: alist2 berghofe@10413: in nipkow@11504: mk_mss (rules,(alist2,weak2),procs,bounds,prems,mk_rews,termless,depth) berghofe@10413: end; berghofe@10413: berghofe@10413: val (op del_congs) = foldl del_cong; berghofe@10413: berghofe@10413: berghofe@10413: (* add_simprocs *) berghofe@10413: nipkow@11504: fun add_proc (mss as Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, berghofe@10413: (name, lhs, proc, id)) = berghofe@10413: let val {sign, t, ...} = rep_cterm lhs berghofe@10413: in (trace_term false ("Adding simplification procedure " ^ quote name ^ " for") berghofe@10413: sign t; berghofe@10413: mk_mss (rules, congs, berghofe@10413: Net.insert_term ((t, mk_simproc (name, proc, lhs, id)), procs, eq_simproc) wenzelm@12603: handle Net.INSERT => wenzelm@12603: (warning ("Ignoring duplicate simplification procedure \"" wenzelm@12603: ^ name ^ "\""); wenzelm@12603: procs), nipkow@11504: bounds, prems, mk_rews, termless,depth)) berghofe@10413: end; berghofe@10413: berghofe@10413: fun add_simproc (mss, (name, lhss, proc, id)) = berghofe@10413: foldl add_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss); berghofe@10413: berghofe@10413: val add_simprocs = foldl add_simproc; berghofe@10413: berghofe@10413: berghofe@10413: (* del_simprocs *) berghofe@10413: nipkow@11504: fun del_proc (mss as Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, berghofe@10413: (name, lhs, proc, id)) = berghofe@10413: mk_mss (rules, congs, berghofe@10413: Net.delete_term ((term_of lhs, mk_simproc (name, proc, lhs, id)), procs, eq_simproc) wenzelm@12603: handle Net.DELETE => wenzelm@12603: (warning ("Simplification procedure \"" ^ name ^ wenzelm@12603: "\" not in simpset"); procs), nipkow@11504: bounds, prems, mk_rews, termless, depth); berghofe@10413: berghofe@10413: fun del_simproc (mss, (name, lhss, proc, id)) = berghofe@10413: foldl del_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss); berghofe@10413: berghofe@10413: val del_simprocs = foldl del_simproc; berghofe@10413: berghofe@10413: berghofe@10413: (* prems *) berghofe@10413: nipkow@11504: fun add_prems (Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, thms) = nipkow@11504: mk_mss (rules, congs, procs, bounds, thms @ prems, mk_rews, termless, depth); berghofe@10413: berghofe@10413: fun prems_of_mss (Mss {prems, ...}) = prems; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_rews *) berghofe@10413: berghofe@10413: fun set_mk_rews nipkow@11504: (Mss {rules, congs, procs, bounds, prems, mk_rews, termless, depth}, mk) = berghofe@10413: mk_mss (rules, congs, procs, bounds, prems, berghofe@10413: {mk=mk, mk_sym= #mk_sym mk_rews, mk_eq_True= #mk_eq_True mk_rews}, nipkow@11504: termless, depth); berghofe@10413: berghofe@10413: fun set_mk_sym nipkow@11504: (Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, mk_sym) = berghofe@10413: mk_mss (rules, congs, procs, bounds, prems, berghofe@10413: {mk= #mk mk_rews, mk_sym= mk_sym, mk_eq_True= #mk_eq_True mk_rews}, nipkow@11504: termless,depth); berghofe@10413: berghofe@10413: fun set_mk_eq_True nipkow@11504: (Mss {rules,congs,procs,bounds,prems,mk_rews,termless,depth}, mk_eq_True) = berghofe@10413: mk_mss (rules, congs, procs, bounds, prems, berghofe@10413: {mk= #mk mk_rews, mk_sym= #mk_sym mk_rews, mk_eq_True= mk_eq_True}, nipkow@11504: termless,depth); berghofe@10413: skalberg@14242: fun get_mk_rews (Mss {mk_rews,...}) = #mk mk_rews skalberg@14242: fun get_mk_sym (Mss {mk_rews,...}) = #mk_sym mk_rews skalberg@14242: fun get_mk_eq_True (Mss {mk_rews,...}) = #mk_eq_True mk_rews skalberg@14242: berghofe@10413: (* termless *) berghofe@10413: berghofe@10413: fun set_termless nipkow@11504: (Mss {rules, congs, procs, bounds, prems, mk_rews, depth, ...}, termless) = nipkow@11504: mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless, depth); berghofe@10413: berghofe@10413: berghofe@10413: 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: berghofe@10413: val dest_eq = Drule.dest_equals o cprop_of; berghofe@10413: val lhs_of = fst o dest_eq; berghofe@10413: val rhs_of = snd o dest_eq; berghofe@10413: berghofe@10413: fun beta_eta_conversion t = berghofe@10413: let val thm = beta_conversion true t; berghofe@10413: in transitive thm (eta_conversion (rhs_of thm)) end; berghofe@10413: berghofe@10413: fun check_conv msg thm thm' = berghofe@10413: let berghofe@10413: val thm'' = transitive thm (transitive berghofe@10413: (symmetric (beta_eta_conversion (lhs_of thm'))) thm') nipkow@13569: in (if msg then trace_thm "SUCCEEDED" thm' else (); Some thm'') end berghofe@10413: handle THM _ => berghofe@10413: let val {sign, prop = _ $ _ $ prop0, ...} = rep_thm thm; berghofe@10413: in nipkow@13569: (trace_thm "Proved wrong thm (Check subgoaler?)" thm'; berghofe@10413: trace_term false "Should have proved:" sign prop0; berghofe@10413: None) berghofe@10413: end; berghofe@10413: berghofe@10413: berghofe@10413: (* mk_procrule *) berghofe@10413: berghofe@10413: fun mk_procrule thm = berghofe@10413: let val (_,prems,lhs,elhs,rhs,_) = decomp_simp thm berghofe@10413: in if rewrite_rule_extra_vars prems lhs rhs berghofe@10413: then (prthm true "Extra vars on rhs:" thm; []) berghofe@13607: else [mk_rrule2{thm=thm, name="", lhs=lhs, elhs=elhs, perm=false}] berghofe@10413: end; berghofe@10413: berghofe@10413: berghofe@10413: (* conversion to apply the meta simpset to a term *) berghofe@10413: berghofe@10413: (* Since the rewriting strategy is bottom-up, we avoid re-normalizing already berghofe@10413: normalized terms by carrying around the rhs of the rewrite rule just berghofe@10413: applied. This is called the `skeleton'. It is decomposed in parallel berghofe@10413: with the term. Once a Var is encountered, the corresponding term is berghofe@10413: already in normal form. berghofe@10413: skel0 is a dummy skeleton that is to enforce complete normalization. berghofe@10413: *) berghofe@10413: val skel0 = Bound 0; berghofe@10413: berghofe@10413: (* Use rhs as skeleton only if the lhs does not contain unnormalized bits. berghofe@10413: The latter may happen iff there are weak congruence rules for constants berghofe@10413: in the lhs. berghofe@10413: *) berghofe@10413: fun uncond_skel((_,weak),(lhs,rhs)) = berghofe@10413: if null weak then rhs (* optimization *) berghofe@10413: else if exists_Const (fn (c,_) => c mem weak) lhs then skel0 berghofe@10413: else rhs; berghofe@10413: berghofe@10413: (* Behaves like unconditional rule if rhs does not contain vars not in the lhs. berghofe@10413: Otherwise those vars may become instantiated with unnormalized terms berghofe@10413: while the premises are solved. berghofe@10413: *) berghofe@10413: fun cond_skel(args as (congs,(lhs,rhs))) = berghofe@10413: if term_varnames rhs subset term_varnames lhs then uncond_skel(args) berghofe@10413: else skel0; berghofe@10413: berghofe@10413: (* berghofe@10413: 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: *) berghofe@10413: berghofe@10413: fun rewritec (prover, signt, maxt) nipkow@11504: (mss as Mss{rules, procs, termless, prems, congs, depth,...}) t = berghofe@10413: let 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 _ = if Sign.subsig (sign, signt) then () berghofe@10413: else (prthm true "Ignoring rewrite rule from different theory:" thm; berghofe@10413: raise Pattern.MATCH); 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); berghofe@10413: val prop' = #prop (rep_thm 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); nipkow@13569: 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; berghofe@10413: val Some thm'' = check_conv false eta_thm thm' berghofe@10413: in Some (thm'', uncond_skel (congs, lr)) end) berghofe@10413: else nipkow@13569: (trace_thm "Trying to rewrite:" thm'; nipkow@13828: case incr_depth mss of nipkow@13828: None => (trace_thm "FAILED - reached depth limit" thm'; None) nipkow@13828: | Some mss => nipkow@13828: (case prover mss thm' of nipkow@13569: None => (trace_thm "FAILED" thm'; None) berghofe@10413: | Some thm2 => berghofe@10413: (case check_conv true eta_thm thm2 of berghofe@10413: None => None | berghofe@10413: Some thm2' => berghofe@10413: let val concl = Logic.strip_imp_concl prop berghofe@10413: val lr = Logic.dest_equals concl nipkow@13828: in Some (thm2', cond_skel (congs, lr)) end)))) berghofe@10413: end berghofe@10413: berghofe@10413: fun rews [] = None berghofe@10413: | rews (rrule :: rrules) = berghofe@10413: let val opt = rew rrule handle Pattern.MATCH => None berghofe@10413: in case opt of None => rews rrules | some => some end; berghofe@10413: berghofe@10413: fun sort_rrules rrs = let wenzelm@12603: fun is_simple({thm, ...}:rrule) = case #prop (rep_thm 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: berghofe@10413: fun proc_rews ([]:simproc list) = None berghofe@10413: | proc_rews ({name, proc, lhs, ...} :: ps) = berghofe@10413: if Pattern.matches tsigt (term_of lhs, 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@13486: (fn () => proc signt prems eta_t) () of wenzelm@13486: None => (debug false "FAILED"; proc_rews ps) wenzelm@13486: | Some raw_thm => nipkow@13569: (trace_thm ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm; berghofe@10413: (case rews (mk_procrule raw_thm) of wenzelm@13486: 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 berghofe@10413: 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 berghofe@10413: 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 = berghofe@10413: let val {sign, ...} = rep_thm cong berghofe@10413: val _ = if Sign.subsig (sign, signt) then () berghofe@10413: else error("Congruence rule from different theory") 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'; ballarin@13932: fun err (msg, thm) = (trace_thm msg thm; None) berghofe@10413: in case prover thm' of ballarin@13932: None => err ("Congruence proof failed. Could not prove", thm') berghofe@10413: | Some thm2 => (case check_conv true (beta_eta_conversion t) thm2 of ballarin@13932: None => err ("Congruence proof failed. Should not have proved", thm2) berghofe@12155: | Some thm2' => berghofe@12155: if op aconv (pairself term_of (dest_equals (cprop_of thm2'))) berghofe@12155: 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: berghofe@13607: fun transitive1 None None = None berghofe@13607: | transitive1 (Some thm1) None = Some thm1 berghofe@13607: | transitive1 None (Some thm2) = Some thm2 berghofe@13607: | transitive1 (Some thm1) (Some thm2) = Some (transitive thm1 thm2) berghofe@10413: berghofe@13607: fun transitive2 thm = transitive1 (Some thm); berghofe@13607: fun transitive3 thm = transitive1 thm o Some; berghofe@13607: berghofe@13607: fun imp_cong' e = combination (combination refl_implies e); berghofe@12155: berghofe@10413: fun bottomc ((simprem,useprem,mutsimp), prover, sign, maxidx) = berghofe@10413: let berghofe@10413: fun botc skel mss t = berghofe@10413: if is_Var skel then None berghofe@10413: else berghofe@10413: (case subc skel mss t of berghofe@10413: some as Some thm1 => berghofe@10413: (case rewritec (prover, sign, maxidx) mss (rhs_of thm1) of berghofe@10413: Some (thm2, skel2) => berghofe@13607: transitive2 (transitive thm1 thm2) berghofe@10413: (botc skel2 mss (rhs_of thm2)) berghofe@10413: | None => some) berghofe@10413: | None => berghofe@10413: (case rewritec (prover, sign, maxidx) mss t of berghofe@13607: Some (thm2, skel2) => transitive2 thm2 berghofe@10413: (botc skel2 mss (rhs_of thm2)) berghofe@10413: | None => None)) berghofe@10413: berghofe@10413: and try_botc mss t = berghofe@10413: (case botc skel0 mss t of berghofe@10413: Some trec1 => trec1 | None => (reflexive t)) berghofe@10413: berghofe@10413: and subc skel nipkow@11504: (mss as Mss{rules,congs,procs,bounds,prems,mk_rews,termless,depth}) t0 = berghofe@10413: (case term_of t0 of berghofe@10413: Abs (a, T, t) => berghofe@10413: let val b = variant bounds a wenzelm@10767: val (v, t') = Thm.dest_abs (Some ("." ^ b)) t0 nipkow@11504: val mss' = mk_mss (rules, congs, procs, b :: bounds, prems, mk_rews, termless,depth) berghofe@10413: val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0 berghofe@10413: in case botc skel' mss' t' of berghofe@10413: Some thm => Some (abstract_rule a v thm) berghofe@10413: | None => None berghofe@10413: end berghofe@10413: | t $ _ => (case t of berghofe@13614: Const ("==>", _) $ _ => impc t0 mss berghofe@10413: | Abs _ => berghofe@10413: let val thm = beta_conversion false t0 berghofe@10413: in case subc skel0 mss (rhs_of thm) of berghofe@10413: None => Some thm berghofe@10413: | 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 berghofe@10413: (case botc tskel mss ct of berghofe@10413: Some thm1 => berghofe@10413: (case botc uskel mss cu of berghofe@10413: Some thm2 => Some (combination thm1 thm2) berghofe@10413: | None => Some (combination thm1 (reflexive cu))) berghofe@10413: | None => berghofe@10413: (case botc uskel mss cu of berghofe@10413: Some thm1 => Some (combination (reflexive ct) thm1) berghofe@10413: | None => None)) berghofe@10413: end berghofe@10413: val (h, ts) = strip_comb t ballarin@13835: in case cong_name h of ballarin@13835: Some a => berghofe@10413: (case assoc_string (fst congs, a) of berghofe@10413: None => appc () berghofe@10413: | Some cong => berghofe@10413: (* post processing: some partial applications h t1 ... tj, j <= length ts, berghofe@10413: may be a redex. Example: map (%x.x) = (%xs.xs) wrt map_cong *) berghofe@10413: (let berghofe@10413: val thm = congc (prover mss, sign, maxidx) cong t0; berghofe@12155: val t = if_none (apsome 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) berghofe@10413: in case botc skel mss cl of berghofe@12155: None => thm berghofe@13607: | 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) berghofe@10413: | _ => None) berghofe@10413: berghofe@13607: and impc ct mss = berghofe@13607: if mutsimp then mut_impc0 [] ct [] [] mss else nonmut_impc ct mss berghofe@10413: berghofe@13607: and rules_of_prem mss prem = berghofe@13607: if maxidx_of_term (term_of prem) <> ~1 berghofe@13607: then (trace_cterm true berghofe@13607: "Cannot add premise as rewrite rule because it contains (type) unknowns:" prem; ([], None)) berghofe@13607: else berghofe@13607: let val asm = assume prem berghofe@13607: in (extract_safe_rrules (mss, asm), Some asm) end berghofe@10413: berghofe@13607: and add_rrules (rrss, asms) mss = berghofe@13607: add_prems (foldl (insert_rrule true) (mss, flat rrss), mapfilter 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 berghofe@13607: | rebuild (prem :: prems) concl (rrs :: rrss) (asm :: asms) mss eq = berghofe@13607: let berghofe@13607: val mss' = add_rrules (rev rrss, rev asms) mss; berghofe@13607: val concl' = berghofe@13607: Drule.mk_implies (prem, if_none (apsome rhs_of eq) concl); berghofe@13607: val dprem = apsome (curry (disch false) prem) berghofe@13607: in case rewritec (prover, sign, maxidx) mss' concl' of berghofe@13607: None => rebuild prems concl' rrss asms mss (dprem eq) berghofe@13607: | Some (eq', _) => transitive2 (foldl (disch false o swap) berghofe@13607: (the (transitive3 (dprem eq) eq'), prems)) berghofe@13607: (mut_impc0 (rev prems) (rhs_of eq') (rev rrss) (rev asms) mss) berghofe@13607: end berghofe@13607: berghofe@13607: and mut_impc0 prems concl rrss asms mss = berghofe@13607: let berghofe@13607: val prems' = strip_imp_prems concl; berghofe@13607: val (rrss', asms') = split_list (map (rules_of_prem mss) prems') berghofe@13607: in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') berghofe@13607: (asms @ asms') [] [] [] [] mss ~1 ~1 berghofe@13607: end berghofe@13607: berghofe@13607: and mut_impc [] concl [] [] prems' rrss' asms' eqns mss changed k = berghofe@13607: transitive1 (foldl (fn (eq2, (eq1, prem)) => transitive1 eq1 berghofe@13607: (apsome (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') berghofe@13607: [] [] [] [] mss ~1 changed berghofe@13607: else rebuild prems' concl rrss' asms' mss berghofe@13607: (botc skel0 (add_rrules (rev rrss', rev asms') mss) concl)) berghofe@13607: berghofe@13607: | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) berghofe@13607: prems' rrss' asms' eqns mss changed k = berghofe@13607: case (if k = 0 then None else botc skel0 (add_rrules berghofe@13607: (rev rrss' @ rrss, rev asms' @ asms) mss) prem) of berghofe@13607: None => mut_impc prems concl rrss asms (prem :: prems') berghofe@13607: (rrs :: rrss') (asm :: asms') (None :: eqns) mss changed berghofe@13607: (if k = 0 then 0 else k - 1) berghofe@13607: | Some eqn => berghofe@13607: let berghofe@13607: val prem' = rhs_of eqn; berghofe@13607: val tprems = map term_of prems; berghofe@13607: val i = 1 + foldl Int.max (~1, map (fn p => berghofe@13607: find_index_eq p tprems) (#hyps (rep_thm eqn))); berghofe@13607: val (rrs', asm') = rules_of_prem mss prem' berghofe@13607: in mut_impc prems concl rrss asms (prem' :: prems') berghofe@13607: (rrs' :: rrss') (asm' :: asms') (Some (foldr (disch true) berghofe@13607: (take (i, prems), imp_cong' eqn (reflexive (Drule.list_implies berghofe@13607: (drop (i, prems), concl))))) :: eqns) mss (length prems') ~1 berghofe@13607: end berghofe@13607: berghofe@10413: (* legacy code - only for backwards compatibility *) berghofe@13607: and nonmut_impc ct mss = berghofe@13607: let val (prem, conc) = dest_implies ct; berghofe@13607: val thm1 = if simprem then botc skel0 mss prem else None; berghofe@10413: val prem1 = if_none (apsome rhs_of thm1) prem; berghofe@13607: val mss1 = if not useprem then mss else add_rrules berghofe@13607: (apsnd single (apfst single (rules_of_prem mss prem1))) mss berghofe@10413: in (case botc skel0 mss1 conc of berghofe@10413: None => (case thm1 of berghofe@10413: None => None berghofe@13607: | Some thm1' => Some (imp_cong' thm1' (reflexive conc))) berghofe@10413: | Some thm2 => berghofe@13607: let val thm2' = disch false (prem1, thm2) berghofe@10413: in (case thm1 of berghofe@10413: None => Some thm2' berghofe@13607: | Some thm1' => berghofe@13607: Some (transitive (imp_cong' thm1' (reflexive conc)) thm2')) berghofe@10413: end) berghofe@10413: end berghofe@10413: berghofe@10413: in try_botc end; berghofe@10413: berghofe@10413: berghofe@10413: (*** 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: mss: contains equality theorems of the form [|p1,...|] ==> t==u berghofe@10413: prover: how to solve premises in conditional rewrites and congruences berghofe@10413: *) berghofe@10413: berghofe@10413: fun rewrite_cterm mode prover mss ct = berghofe@10413: let val {sign, t, maxidx, ...} = rep_cterm ct nipkow@11505: val Mss{depth, ...} = mss nipkow@14330: in trace_cterm false "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" ct; nipkow@14330: simp_depth := depth; nipkow@11505: bottomc (mode, prover, sign, maxidx) mss ct nipkow@11505: end berghofe@10413: 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: berghofe@10413: (*In [A1,...,An]==>B, rewrite the selected A's only -- for rewrite_goals_tac*) berghofe@10413: fun goals_conv pred cv = berghofe@10413: let fun gconv i ct = berghofe@10413: let val (A,B) = Drule.dest_implies ct berghofe@13661: in imp_cong' (if pred i then cv A else reflexive A) (gconv (i+1) B) end berghofe@10413: handle TERM _ => reflexive ct berghofe@10413: in gconv 1 end; berghofe@10413: berghofe@11737: (* Rewrite A in !!x1,...,xn. A *) berghofe@11736: fun forall_conv cv ct = berghofe@11736: let val p as (ct1, ct2) = Thm.dest_comb ct berghofe@11736: in (case pairself term_of p of berghofe@11736: (Const ("all", _), Abs (s, _, _)) => berghofe@11736: let val (v, ct') = Thm.dest_abs (Some "@") ct2; berghofe@11736: in Thm.combination (Thm.reflexive ct1) berghofe@11736: (Thm.abstract_rule s v (forall_conv cv ct')) berghofe@11736: end berghofe@11736: | _ => cv ct) berghofe@11736: end handle TERM _ => cv ct; berghofe@11736: berghofe@10413: (*Use a conversion to transform a theorem*) berghofe@10413: fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th; berghofe@10413: wenzelm@11760: (*Rewrite a cterm*) wenzelm@11767: fun rewrite_aux _ _ [] = (fn ct => Thm.reflexive ct) wenzelm@11767: | rewrite_aux prover full thms = rewrite_cterm (full, false, false) prover (mss_of thms); wenzelm@11672: berghofe@10413: (*Rewrite a theorem*) wenzelm@11767: fun simplify_aux _ _ [] = (fn th => th) wenzelm@11767: | simplify_aux prover full thms = wenzelm@11767: fconv_rule (rewrite_cterm (full, false, false) prover (mss_of thms)); berghofe@10413: berghofe@10413: fun rewrite_thm mode prover mss = fconv_rule (rewrite_cterm mode prover mss); 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 = berghofe@10413: fconv_rule (goals_conv (K true) (rewrite_cterm (true, true, false) prover berghofe@10413: (mss_of thms))) th; berghofe@10413: berghofe@10413: (*Rewrite the subgoal of a proof state (represented by a theorem) *) berghofe@10413: fun rewrite_goal_rule mode prover mss i thm = berghofe@10413: if 0 < i andalso i <= nprems_of thm berghofe@10413: then fconv_rule (goals_conv (fn j => j=i) (rewrite_cterm mode prover mss)) thm berghofe@10413: else raise THM("rewrite_goal_rule",i,[thm]); berghofe@10413: wenzelm@12783: wenzelm@12783: (*simple term rewriting -- without proofs*) berghofe@13196: fun rewrite_term sg rules procs = berghofe@13196: Pattern.rewrite_term (Sign.tsig_of sg) (map decomp_simp' rules) procs; wenzelm@12783: berghofe@10413: end; berghofe@10413: wenzelm@11672: structure BasicMetaSimplifier: BASIC_META_SIMPLIFIER = MetaSimplifier; wenzelm@11672: open BasicMetaSimplifier;