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