src/Pure/simplifier.ML
author wenzelm
Sat, 14 Apr 2012 17:15:57 +0200
changeset 48341 402b753d8383
parent 48110 0b1829860149
child 49791 37cd53e69840
permissions -rw-r--r--
outermost SELECT_GOAL potentially improves performance;
     1 (*  Title:      Pure/simplifier.ML
     2     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
     3 
     4 Generic simplifier, suitable for most logics (see also
     5 raw_simplifier.ML for the actual meta-level rewriting engine).
     6 *)
     7 
     8 signature BASIC_SIMPLIFIER =
     9 sig
    10   include BASIC_RAW_SIMPLIFIER
    11   val map_simpset: (simpset -> simpset) -> Proof.context -> Proof.context
    12   val simpset_of: Proof.context -> simpset
    13   val global_simpset_of: theory -> simpset
    14   val Addsimprocs: simproc list -> unit
    15   val Delsimprocs: simproc list -> unit
    16   val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
    17   val safe_asm_full_simp_tac: simpset -> int -> tactic
    18   val simp_tac: simpset -> int -> tactic
    19   val asm_simp_tac: simpset -> int -> tactic
    20   val full_simp_tac: simpset -> int -> tactic
    21   val asm_lr_simp_tac: simpset -> int -> tactic
    22   val asm_full_simp_tac: simpset -> int -> tactic
    23   val simplify: simpset -> thm -> thm
    24   val asm_simplify: simpset -> thm -> thm
    25   val full_simplify: simpset -> thm -> thm
    26   val asm_lr_simplify: simpset -> thm -> thm
    27   val asm_full_simplify: simpset -> thm -> thm
    28 end;
    29 
    30 signature SIMPLIFIER =
    31 sig
    32   include BASIC_SIMPLIFIER
    33   val map_simpset_global: (simpset -> simpset) -> theory -> theory
    34   val pretty_ss: Proof.context -> simpset -> Pretty.T
    35   val clear_ss: simpset -> simpset
    36   val default_mk_sym: simpset -> thm -> thm option
    37   val debug_bounds: bool Unsynchronized.ref
    38   val prems_of: simpset -> thm list
    39   val add_simp: thm -> simpset -> simpset
    40   val del_simp: thm -> simpset -> simpset
    41   val add_eqcong: thm -> simpset -> simpset
    42   val del_eqcong: thm -> simpset -> simpset
    43   val add_cong: thm -> simpset -> simpset
    44   val del_cong: thm -> simpset -> simpset
    45   val add_prems: thm list -> simpset -> simpset
    46   val mksimps: simpset -> thm -> thm list
    47   val set_mksimps: (simpset -> thm -> thm list) -> simpset -> simpset
    48   val set_mkcong: (simpset -> thm -> thm) -> simpset -> simpset
    49   val set_mksym: (simpset -> thm -> thm option) -> simpset -> simpset
    50   val set_mkeqTrue: (simpset -> thm -> thm option) -> simpset -> simpset
    51   val set_termless: (term * term -> bool) -> simpset -> simpset
    52   val set_subgoaler: (simpset -> int -> tactic) -> simpset -> simpset
    53   val inherit_context: simpset -> simpset -> simpset
    54   val the_context: simpset -> Proof.context
    55   val context: Proof.context -> simpset -> simpset
    56   val global_context: theory -> simpset -> simpset
    57   val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset
    58   val simproc_global_i: theory -> string -> term list ->
    59     (theory -> simpset -> term -> thm option) -> simproc
    60   val simproc_global: theory -> string -> string list ->
    61     (theory -> simpset -> term -> thm option) -> simproc
    62   val rewrite: simpset -> conv
    63   val asm_rewrite: simpset -> conv
    64   val full_rewrite: simpset -> conv
    65   val asm_lr_rewrite: simpset -> conv
    66   val asm_full_rewrite: simpset -> conv
    67   val get_ss: Context.generic -> simpset
    68   val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic
    69   val attrib: (thm -> simpset -> simpset) -> attribute
    70   val simp_add: attribute
    71   val simp_del: attribute
    72   val cong_add: attribute
    73   val cong_del: attribute
    74   val check_simproc: Proof.context -> xstring * Position.T -> string
    75   val the_simproc: Proof.context -> string -> simproc
    76   val def_simproc: {name: binding, lhss: term list,
    77     proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    78     local_theory -> local_theory
    79   val def_simproc_cmd: {name: binding, lhss: string list,
    80     proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    81     local_theory -> local_theory
    82   val cong_modifiers: Method.modifier parser list
    83   val simp_modifiers': Method.modifier parser list
    84   val simp_modifiers: Method.modifier parser list
    85   val method_setup: Method.modifier parser list -> theory -> theory
    86   val easy_setup: thm -> thm list -> theory -> theory
    87 end;
    88 
    89 structure Simplifier: SIMPLIFIER =
    90 struct
    91 
    92 open Raw_Simplifier;
    93 
    94 
    95 (** data **)
    96 
    97 structure Data = Generic_Data
    98 (
    99   type T = simpset * simproc Name_Space.table;
   100   val empty : T = (empty_ss, Name_Space.empty_table "simproc");
   101   fun extend (ss, tab) = (Raw_Simplifier.inherit_context empty_ss ss, tab);
   102   fun merge ((ss1, tab1), (ss2, tab2)) =
   103     (merge_ss (ss1, ss2), Name_Space.merge_tables (tab1, tab2));
   104 );
   105 
   106 val get_ss = fst o Data.get;
   107 
   108 fun map_ss f context =
   109   Data.map (apfst ((Raw_Simplifier.with_context (Context.proof_of context) f))) context;
   110 
   111 val get_simprocs = snd o Data.get o Context.Proof;
   112 
   113 
   114 
   115 (** pretty printing **)
   116 
   117 fun pretty_ss ctxt ss =
   118   let
   119     val pretty_cterm = Syntax.pretty_term ctxt o Thm.term_of;
   120     val pretty_thm = Display.pretty_thm ctxt;
   121     fun pretty_proc (name, lhss) = Pretty.big_list (name ^ ":") (map pretty_cterm lhss);
   122     fun pretty_cong (name, thm) =
   123       Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, pretty_thm thm];
   124 
   125     val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss;
   126   in
   127     [Pretty.big_list "simplification rules:" (map (pretty_thm o #2) simps),
   128       Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)),
   129       Pretty.big_list "congruences:" (map pretty_cong congs),
   130       Pretty.strs ("loopers:" :: map quote loopers),
   131       Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
   132       Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
   133     |> Pretty.chunks
   134   end;
   135 
   136 
   137 
   138 (** simpset data **)
   139 
   140 (* attributes *)
   141 
   142 fun attrib f = Thm.declaration_attribute (map_ss o f);
   143 
   144 val simp_add = attrib add_simp;
   145 val simp_del = attrib del_simp;
   146 val cong_add = attrib add_cong;
   147 val cong_del = attrib del_cong;
   148 
   149 
   150 (* local simpset *)
   151 
   152 fun map_simpset f = Context.proof_map (map_ss f);
   153 fun simpset_of ctxt = Raw_Simplifier.context ctxt (get_ss (Context.Proof ctxt));
   154 
   155 val _ = Context.>> (Context.map_theory
   156   (ML_Antiquote.value (Binding.name "simpset")
   157     (Scan.succeed "Simplifier.simpset_of (ML_Context.the_local_context ())")));
   158 
   159 
   160 (* global simpset *)
   161 
   162 fun map_simpset_global f = Context.theory_map (map_ss f);
   163 fun global_simpset_of thy =
   164   Raw_Simplifier.context (Proof_Context.init_global thy) (get_ss (Context.Theory thy));
   165 
   166 fun Addsimprocs args = Context.>> (map_ss (fn ss => ss addsimprocs args));
   167 fun Delsimprocs args = Context.>> (map_ss (fn ss => ss delsimprocs args));
   168 
   169 
   170 
   171 (** named simprocs **)
   172 
   173 (* get simprocs *)
   174 
   175 fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1;
   176 val the_simproc = Name_Space.get o get_simprocs;
   177 
   178 val _ =
   179   Context.>> (Context.map_theory
   180    (ML_Antiquote.value (Binding.name "simproc")
   181       (Args.context -- Scan.lift (Parse.position Args.name)
   182         >> (fn (ctxt, name) =>
   183           "Simplifier.the_simproc (ML_Context.the_local_context ()) " ^
   184             ML_Syntax.print_string (check_simproc ctxt name)))));
   185 
   186 
   187 (* define simprocs *)
   188 
   189 local
   190 
   191 fun gen_simproc prep {name = b, lhss, proc, identifier} lthy =
   192   let
   193     val simproc = make_simproc
   194       {name = Local_Theory.full_name lthy b,
   195        lhss =
   196         let
   197           val lhss' = prep lthy lhss;
   198           val ctxt' = fold Variable.auto_fixes lhss' lthy;
   199         in Variable.export_terms ctxt' lthy lhss' end
   200         |> map (Thm.cterm_of (Proof_Context.theory_of lthy)),
   201        proc = proc,
   202        identifier = identifier};
   203   in
   204     lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context =>
   205       let
   206         val b' = Morphism.binding phi b;
   207         val simproc' = transform_simproc phi simproc;
   208       in
   209         context
   210         |> Data.map (fn (ss, tab) =>
   211           (ss addsimprocs [simproc'], #2 (Name_Space.define context true (b', simproc') tab)))
   212       end)
   213   end;
   214 
   215 in
   216 
   217 val def_simproc = gen_simproc Syntax.check_terms;
   218 val def_simproc_cmd = gen_simproc Syntax.read_terms;
   219 
   220 end;
   221 
   222 
   223 
   224 (** simplification tactics and rules **)
   225 
   226 fun solve_all_tac solvers ss =
   227   let
   228     val (_, {subgoal_tac, ...}) = Raw_Simplifier.internal_ss ss;
   229     val solve_tac = subgoal_tac (Raw_Simplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac);
   230   in DEPTH_SOLVE (solve_tac 1) end;
   231 
   232 (*NOTE: may instantiate unknowns that appear also in other subgoals*)
   233 fun generic_simp_tac safe mode ss =
   234   let
   235     val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = Raw_Simplifier.internal_ss ss;
   236     val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs));
   237     val solve_tac = FIRST' (map (Raw_Simplifier.solver ss)
   238       (rev (if safe then solvers else unsafe_solvers)));
   239 
   240     fun simp_loop_tac i =
   241       Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN
   242       (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));
   243   in SELECT_GOAL (simp_loop_tac 1) end;
   244 
   245 local
   246 
   247 fun simp rew mode ss thm =
   248   let
   249     val (_, {solvers = (unsafe_solvers, _), ...}) = Raw_Simplifier.internal_ss ss;
   250     val tacf = solve_all_tac (rev unsafe_solvers);
   251     fun prover s th = Option.map #1 (Seq.pull (tacf s th));
   252   in rew mode prover ss thm end;
   253 
   254 in
   255 
   256 val simp_thm = simp Raw_Simplifier.rewrite_thm;
   257 val simp_cterm = simp Raw_Simplifier.rewrite_cterm;
   258 
   259 end;
   260 
   261 
   262 (* tactics *)
   263 
   264 val simp_tac = generic_simp_tac false (false, false, false);
   265 val asm_simp_tac = generic_simp_tac false (false, true, false);
   266 val full_simp_tac = generic_simp_tac false (true, false, false);
   267 val asm_lr_simp_tac = generic_simp_tac false (true, true, false);
   268 val asm_full_simp_tac = generic_simp_tac false (true, true, true);
   269 val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);
   270 
   271 
   272 (* conversions *)
   273 
   274 val          simplify = simp_thm (false, false, false);
   275 val      asm_simplify = simp_thm (false, true, false);
   276 val     full_simplify = simp_thm (true, false, false);
   277 val   asm_lr_simplify = simp_thm (true, true, false);
   278 val asm_full_simplify = simp_thm (true, true, true);
   279 
   280 val          rewrite = simp_cterm (false, false, false);
   281 val      asm_rewrite = simp_cterm (false, true, false);
   282 val     full_rewrite = simp_cterm (true, false, false);
   283 val   asm_lr_rewrite = simp_cterm (true, true, false);
   284 val asm_full_rewrite = simp_cterm (true, true, true);
   285 
   286 
   287 
   288 (** concrete syntax of attributes **)
   289 
   290 (* add / del *)
   291 
   292 val simpN = "simp";
   293 val congN = "cong";
   294 val onlyN = "only";
   295 val no_asmN = "no_asm";
   296 val no_asm_useN = "no_asm_use";
   297 val no_asm_simpN = "no_asm_simp";
   298 val asm_lrN = "asm_lr";
   299 
   300 
   301 (* simprocs *)
   302 
   303 local
   304 
   305 val add_del =
   306   (Args.del -- Args.colon >> K (op delsimprocs) ||
   307     Scan.option (Args.add -- Args.colon) >> K (op addsimprocs))
   308   >> (fn f => fn simproc => fn phi => Thm.declaration_attribute
   309       (K (map_ss (fn ss => f (ss, [transform_simproc phi simproc])))));
   310 
   311 in
   312 
   313 val simproc_att =
   314   (Args.context -- Scan.lift add_del) :|-- (fn (ctxt, decl) =>
   315     Scan.repeat1 (Scan.lift (Args.named_attribute (decl o the_simproc ctxt o check_simproc ctxt))))
   316   >> (fn atts => Thm.declaration_attribute (fn th =>
   317         fold (fn att => Thm.attribute_declaration (Morphism.form att) th) atts));
   318 
   319 end;
   320 
   321 
   322 (* conversions *)
   323 
   324 local
   325 
   326 fun conv_mode x =
   327   ((Args.parens (Args.$$$ no_asmN) >> K simplify ||
   328     Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify ||
   329     Args.parens (Args.$$$ no_asm_useN) >> K full_simplify ||
   330     Scan.succeed asm_full_simplify) |> Scan.lift) x;
   331 
   332 in
   333 
   334 val simplified = conv_mode -- Attrib.thms >>
   335   (fn (f, ths) => Thm.rule_attribute (fn context =>
   336     f ((if null ths then I else Raw_Simplifier.clear_ss)
   337         (simpset_of (Context.proof_of context)) addsimps ths)));
   338 
   339 end;
   340 
   341 
   342 (* setup attributes *)
   343 
   344 val _ = Context.>> (Context.map_theory
   345  (Attrib.setup (Binding.name simpN) (Attrib.add_del simp_add simp_del)
   346     "declaration of Simplifier rewrite rule" #>
   347   Attrib.setup (Binding.name congN) (Attrib.add_del cong_add cong_del)
   348     "declaration of Simplifier congruence rule" #>
   349   Attrib.setup (Binding.name "simproc") simproc_att
   350     "declaration of simplification procedures" #>
   351   Attrib.setup (Binding.name "simplified") simplified "simplified rule"));
   352 
   353 
   354 
   355 (** method syntax **)
   356 
   357 val cong_modifiers =
   358  [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),
   359   Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),
   360   Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];
   361 
   362 val simp_modifiers =
   363  [Args.$$$ simpN -- Args.colon >> K (I, simp_add),
   364   Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),
   365   Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),
   366   Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon
   367     >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)]
   368    @ cong_modifiers;
   369 
   370 val simp_modifiers' =
   371  [Args.add -- Args.colon >> K (I, simp_add),
   372   Args.del -- Args.colon >> K (I, simp_del),
   373   Args.$$$ onlyN -- Args.colon
   374     >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)]
   375    @ cong_modifiers;
   376 
   377 val simp_options =
   378  (Args.parens (Args.$$$ no_asmN) >> K simp_tac ||
   379   Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac ||
   380   Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac ||
   381   Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac ||
   382   Scan.succeed asm_full_simp_tac);
   383 
   384 fun simp_method more_mods meth =
   385   Scan.lift simp_options --|
   386     Method.sections (more_mods @ simp_modifiers') >>
   387     (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts));
   388 
   389 
   390 
   391 (** setup **)
   392 
   393 fun method_setup more_mods =
   394   Method.setup (Binding.name simpN)
   395     (simp_method more_mods (fn ctxt => fn tac => fn facts =>
   396       HEADGOAL (Method.insert_tac facts THEN'
   397         (CHANGED_PROP oo tac) (simpset_of ctxt))))
   398     "simplification" #>
   399   Method.setup (Binding.name "simp_all")
   400     (simp_method more_mods (fn ctxt => fn tac => fn facts =>
   401       ALLGOALS (Method.insert_tac facts) THEN
   402         (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) (simpset_of ctxt)))
   403     "simplification (all goals)";
   404 
   405 fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ =>
   406   let
   407     val trivialities = Drule.reflexive_thm :: trivs;
   408 
   409     fun unsafe_solver_tac ss =
   410       FIRST' [resolve_tac (trivialities @ Raw_Simplifier.prems_of ss), assume_tac];
   411     val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;
   412 
   413     (*no premature instantiation of variables during simplification*)
   414     fun safe_solver_tac ss =
   415       FIRST' [match_tac (trivialities @ Raw_Simplifier.prems_of ss), eq_assume_tac];
   416     val safe_solver = mk_solver "easy safe" safe_solver_tac;
   417 
   418     fun mk_eq thm =
   419       if can Logic.dest_equals (Thm.concl_of thm) then [thm]
   420       else [thm RS reflect] handle THM _ => [];
   421 
   422     fun mksimps thm = mk_eq (Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm);
   423   in
   424     empty_ss
   425     setSSolver safe_solver
   426     setSolver unsafe_solver
   427     |> set_subgoaler asm_simp_tac
   428     |> set_mksimps (K mksimps)
   429   end));
   430 
   431 end;
   432 
   433 structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier;
   434 open Basic_Simplifier;