src/Pure/simplifier.ML
author wenzelm
Thu, 27 Mar 2008 15:32:15 +0100
changeset 26435 bdce320cd426
parent 26425 6561665c5cb1
child 26463 9283b4185fdf
permissions -rw-r--r--
eliminated delayed theory setup

(*  Title:      Pure/simplifier.ML
    ID:         $Id$
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen

Generic simplifier, suitable for most logics (see also
meta_simplifier.ML for the actual meta-level rewriting engine).
*)

signature BASIC_SIMPLIFIER =
sig
  include BASIC_META_SIMPLIFIER
  val print_simpset: theory -> unit
  val change_simpset_of: theory -> (simpset -> simpset) -> unit
  val change_simpset: (simpset -> simpset) -> unit
  val simpset_of: theory -> simpset
  val simpset: unit -> simpset
  val SIMPSET: (simpset -> tactic) -> tactic
  val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic
  val Addsimps: thm list -> unit
  val Delsimps: thm list -> unit
  val Addsimprocs: simproc list -> unit
  val Delsimprocs: simproc list -> unit
  val Addcongs: thm list -> unit
  val Delcongs: thm list -> unit
  val local_simpset_of: Proof.context -> simpset
  val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
  val safe_asm_full_simp_tac: simpset -> int -> tactic
  val               simp_tac: simpset -> int -> tactic
  val           asm_simp_tac: simpset -> int -> tactic
  val          full_simp_tac: simpset -> int -> tactic
  val        asm_lr_simp_tac: simpset -> int -> tactic
  val      asm_full_simp_tac: simpset -> int -> tactic
  val               Simp_tac:            int -> tactic
  val           Asm_simp_tac:            int -> tactic
  val          Full_simp_tac:            int -> tactic
  val        Asm_lr_simp_tac:            int -> tactic
  val      Asm_full_simp_tac:            int -> tactic
  val          simplify: simpset -> thm -> thm
  val      asm_simplify: simpset -> thm -> thm
  val     full_simplify: simpset -> thm -> thm
  val   asm_lr_simplify: simpset -> thm -> thm
  val asm_full_simplify: simpset -> thm -> thm
end;

signature SIMPLIFIER =
sig
  include BASIC_SIMPLIFIER
  val clear_ss: simpset -> simpset
  val debug_bounds: bool ref
  val inherit_context: simpset -> simpset -> simpset
  val the_context: simpset -> Proof.context
  val context: Proof.context -> simpset -> simpset
  val theory_context: theory  -> simpset -> simpset
  val simproc_i: theory -> string -> term list
    -> (theory -> simpset -> term -> thm option) -> simproc
  val simproc: theory -> string -> string list
    -> (theory -> simpset -> term -> thm option) -> simproc
  val          rewrite: simpset -> conv
  val      asm_rewrite: simpset -> conv
  val     full_rewrite: simpset -> conv
  val   asm_lr_rewrite: simpset -> conv
  val asm_full_rewrite: simpset -> conv
  val get_simpset: theory -> simpset
  val print_local_simpset: Proof.context -> unit
  val get_local_simpset: Proof.context -> simpset
  val put_local_simpset: simpset -> Proof.context -> Proof.context
  val get_ss: Context.generic -> simpset
  val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic
  val attrib: (simpset * thm list -> simpset) -> attribute
  val simp_add: attribute
  val simp_del: attribute
  val cong_add: attribute
  val cong_del: attribute
  val get_simproc: Context.generic -> xstring -> simproc
  val def_simproc: {name: string, lhss: string list,
    proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    local_theory -> local_theory
  val def_simproc_i: {name: string, lhss: term list,
    proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    local_theory -> local_theory
  val cong_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
  val simp_modifiers': (Args.T list -> (Method.modifier * Args.T list)) list
  val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
  val method_setup: (Args.T list -> (Method.modifier * Args.T list)) list
    -> theory -> theory
  val easy_setup: thm -> thm list -> theory -> theory
end;

structure Simplifier: SIMPLIFIER =
struct

open MetaSimplifier;


(** simpset data **)

(* global simpsets *)

structure GlobalSimpset = TheoryDataFun
(
  type T = simpset ref;
  val empty = ref empty_ss;
  fun copy (ref ss) = ref ss: T;
  fun extend (ref ss) = ref (MetaSimplifier.inherit_context empty_ss ss);
  fun merge _ (ref ss1, ref ss2) = ref (merge_ss (ss1, ss2));
);

val _ = Context.>> GlobalSimpset.init;
fun print_simpset thy = print_ss (! (GlobalSimpset.get thy));
val get_simpset = ! o GlobalSimpset.get;

fun change_simpset_of thy f = CRITICAL (fn () => change (GlobalSimpset.get thy) f);
fun change_simpset f = CRITICAL (fn () => change_simpset_of (ML_Context.the_global_context ()) f);

fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_simpset thy);
val simpset = simpset_of o ML_Context.the_global_context;


fun SIMPSET tacf st = tacf (simpset_of (Thm.theory_of_thm st)) st;
fun SIMPSET' tacf i st = tacf (simpset_of (Thm.theory_of_thm st)) i st;

fun Addsimps args = change_simpset (fn ss => ss addsimps args);
fun Delsimps args = change_simpset (fn ss => ss delsimps args);
fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args);
fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args);
fun Addcongs args = change_simpset (fn ss => ss addcongs args);
fun Delcongs args = change_simpset (fn ss => ss delcongs args);


(* local simpsets *)

structure LocalSimpset = ProofDataFun
(
  type T = simpset;
  val init = get_simpset;
);

val print_local_simpset = print_ss o LocalSimpset.get;
val get_local_simpset = LocalSimpset.get;
val put_local_simpset = LocalSimpset.put;

fun local_simpset_of ctxt = MetaSimplifier.context ctxt (get_local_simpset ctxt);

val _ = ML_Context.value_antiq "simpset"
  (Scan.succeed ("simpset", "Simplifier.local_simpset_of (ML_Context.the_local_context ())"));


(* generic simpsets *)

fun get_ss (Context.Theory thy) = simpset_of thy
  | get_ss (Context.Proof ctxt) = local_simpset_of ctxt;

fun map_ss f (Context.Theory thy) = (change_simpset_of thy f; Context.Theory thy)
  | map_ss f (Context.Proof ctxt) = Context.Proof (LocalSimpset.map f ctxt);


(* attributes *)

fun attrib f = Thm.declaration_attribute (fn th => map_ss (fn ss => f (ss, [th])));

val simp_add = attrib (op addsimps);
val simp_del = attrib (op delsimps);
val cong_add = attrib (op addcongs);
val cong_del = attrib (op delcongs);



(** named simprocs **)

fun err_dup_simproc name = error ("Duplicate simproc: " ^ quote name);


(* data *)

structure Simprocs = GenericDataFun
(
  type T = simproc NameSpace.table;
  val empty = NameSpace.empty_table;
  val extend = I;
  fun merge _ simprocs = NameSpace.merge_tables eq_simproc simprocs
    handle Symtab.DUP dup => err_dup_simproc dup;
);


(* get simprocs *)

fun get_simproc context xname =
  let
    val (space, tab) = Simprocs.get context;
    val name = NameSpace.intern space xname;
  in
    (case Symtab.lookup tab name of
      SOME proc => proc
    | NONE => error ("Undefined simplification procedure: " ^ quote name))
  end;

val _ = ML_Context.value_antiq "simproc" (Scan.lift Args.name >> (fn name =>
  ("simproc",
    "Simplifier.get_simproc (ML_Context.the_generic_context ()) " ^ ML_Syntax.print_string name)));


(* define simprocs *)

local

fun gen_simproc prep {name, lhss, proc, identifier} lthy =
  let
    val naming = LocalTheory.full_naming lthy;
    val simproc = make_simproc
      {name = LocalTheory.full_name lthy name,
       lhss =
        let
          val lhss' = prep lthy lhss;
          val ctxt' = lthy
            |> fold Variable.declare_term lhss'
            |> fold Variable.auto_fixes lhss';
        in Variable.export_terms ctxt' lthy lhss' end
        |> map (Thm.cterm_of (ProofContext.theory_of lthy)),
       proc = proc,
       identifier = identifier}
      |> morph_simproc (LocalTheory.target_morphism lthy);
  in
    lthy |> LocalTheory.declaration (fn phi =>
      let
        val name' = Morphism.name phi name;
        val simproc' = morph_simproc phi simproc;
      in
        Simprocs.map (fn simprocs =>
            NameSpace.extend_table naming [(name', simproc')] simprocs
              handle Symtab.DUP dup => err_dup_simproc dup)
        #> map_ss (fn ss => ss addsimprocs [simproc'])
      end)
  end;

in

val def_simproc = gen_simproc Syntax.read_terms;
val def_simproc_i = gen_simproc Syntax.check_terms;

end;



(** simplification tactics and rules **)

fun solve_all_tac solvers ss =
  let
    val (_, {subgoal_tac, ...}) = MetaSimplifier.rep_ss ss;
    val solve_tac = subgoal_tac (MetaSimplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac);
  in DEPTH_SOLVE (solve_tac 1) end;

(*NOTE: may instantiate unknowns that appear also in other subgoals*)
fun generic_simp_tac safe mode ss =
  let
    val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = MetaSimplifier.rep_ss ss;
    val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs));
    val solve_tac = FIRST' (map (MetaSimplifier.solver ss)
      (rev (if safe then solvers else unsafe_solvers)));

    fun simp_loop_tac i =
      asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN
      (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));
  in simp_loop_tac end;

local

fun simp rew mode ss thm =
  let
    val (_, {solvers = (unsafe_solvers, _), ...}) = MetaSimplifier.rep_ss ss;
    val tacf = solve_all_tac (rev unsafe_solvers);
    fun prover s th = Option.map #1 (Seq.pull (tacf s th));
  in rew mode prover ss thm end;

in

val simp_thm = simp MetaSimplifier.rewrite_thm;
val simp_cterm = simp MetaSimplifier.rewrite_cterm;

end;


(* tactics *)

val simp_tac = generic_simp_tac false (false, false, false);
val asm_simp_tac = generic_simp_tac false (false, true, false);
val full_simp_tac = generic_simp_tac false (true, false, false);
val asm_lr_simp_tac = generic_simp_tac false (true, true, false);
val asm_full_simp_tac = generic_simp_tac false (true, true, true);
val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);

(*the abstraction over the proof state delays the dereferencing*)
fun          Simp_tac i st =          simp_tac (simpset ()) i st;
fun      Asm_simp_tac i st =      asm_simp_tac (simpset ()) i st;
fun     Full_simp_tac i st =     full_simp_tac (simpset ()) i st;
fun   Asm_lr_simp_tac i st =   asm_lr_simp_tac (simpset ()) i st;
fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st;


(* conversions *)

val          simplify = simp_thm (false, false, false);
val      asm_simplify = simp_thm (false, true, false);
val     full_simplify = simp_thm (true, false, false);
val   asm_lr_simplify = simp_thm (true, true, false);
val asm_full_simplify = simp_thm (true, true, true);

val          rewrite = simp_cterm (false, false, false);
val      asm_rewrite = simp_cterm (false, true, false);
val     full_rewrite = simp_cterm (true, false, false);
val   asm_lr_rewrite = simp_cterm (true, true, false);
val asm_full_rewrite = simp_cterm (true, true, true);



(** concrete syntax of attributes **)

(* add / del *)

val simpN = "simp";
val congN = "cong";
val addN = "add";
val delN = "del";
val onlyN = "only";
val no_asmN = "no_asm";
val no_asm_useN = "no_asm_use";
val no_asm_simpN = "no_asm_simp";
val asm_lrN = "asm_lr";


(* simprocs *)

local

val add_del =
  (Args.del -- Args.colon >> K (op delsimprocs) ||
    Scan.option (Args.add -- Args.colon) >> K (op addsimprocs))
  >> (fn f => fn simproc => fn phi => Thm.declaration_attribute
      (K (map_ss (fn ss => f (ss, [morph_simproc phi simproc])))));

in

val simproc_att = Attrib.syntax
  (Scan.peek (fn context =>
    add_del :|-- (fn decl =>
      Scan.repeat1 (Args.named_attribute (decl o get_simproc context))
      >> (Library.apply o map Morphism.form))));

end;


(* conversions *)

local

fun conv_mode x =
  ((Args.parens (Args.$$$ no_asmN) >> K simplify ||
    Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify ||
    Args.parens (Args.$$$ no_asm_useN) >> K full_simplify ||
    Scan.succeed asm_full_simplify) |> Scan.lift) x;

in

val simplified =
  Attrib.syntax (conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute (fn x =>
    f ((if null ths then I else MetaSimplifier.clear_ss) (get_ss x) addsimps ths))));

end;


(* setup attributes *)

val _ = Context.>>
 (Attrib.add_attributes
   [(simpN, Attrib.add_del_args simp_add simp_del, "declaration of Simplifier rewrite rule"),
    (congN, Attrib.add_del_args cong_add cong_del, "declaration of Simplifier congruence rule"),
    ("simproc", simproc_att, "declaration of simplification procedures"),
    ("simplified", simplified, "simplified rule")]);



(** proof methods **)

(* simplification *)

val simp_options =
 (Args.parens (Args.$$$ no_asmN) >> K simp_tac ||
  Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac ||
  Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac ||
  Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac ||
  Scan.succeed asm_full_simp_tac);

val cong_modifiers =
 [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),
  Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),
  Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];

val simp_modifiers =
 [Args.$$$ simpN -- Args.colon >> K (I, simp_add),
  Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),
  Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),
  Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon
    >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add)]
   @ cong_modifiers;

val simp_modifiers' =
 [Args.add -- Args.colon >> K (I, simp_add),
  Args.del -- Args.colon >> K (I, simp_del),
  Args.$$$ onlyN -- Args.colon
    >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add)]
   @ cong_modifiers;

fun simp_args more_mods =
  Method.sectioned_args (Args.bang_facts -- Scan.lift simp_options)
    (more_mods @ simp_modifiers');

fun simp_method (prems, tac) ctxt = Method.METHOD (fn facts =>
  ALLGOALS (Method.insert_tac (prems @ facts)) THEN
    (CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt));

fun simp_method' (prems, tac) ctxt = Method.METHOD (fn facts =>
  HEADGOAL (Method.insert_tac (prems @ facts) THEN'
      ((CHANGED_PROP) oo tac) (local_simpset_of ctxt)));



(** setup **)

fun method_setup more_mods = Method.add_methods
 [(simpN, simp_args more_mods simp_method', "simplification"),
  ("simp_all", simp_args more_mods simp_method, "simplification (all goals)")];

fun easy_setup reflect trivs = method_setup [] #> (fn thy =>
  let
    val trivialities = Drule.reflexive_thm :: trivs;

    fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac];
    val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;

    (*no premature instantiation of variables during simplification*)
    fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac];
    val safe_solver = mk_solver "easy safe" safe_solver_tac;

    fun mk_eq thm =
      if can Logic.dest_equals (Thm.concl_of thm) then [thm]
      else [thm RS reflect] handle THM _ => [];

    fun mksimps thm = mk_eq (Drule.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm);
    val _ = CRITICAL (fn () =>
      GlobalSimpset.get thy :=
        empty_ss setsubgoaler asm_simp_tac
        setSSolver safe_solver
        setSolver unsafe_solver
        setmksimps mksimps);
  in thy end);

end;

structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier;
open BasicSimplifier;