src/Pure/simplifier.ML
author wenzelm
Sun, 08 Mar 2009 12:15:58 +0100
changeset 30356 36d0e00af606
parent 30342 efd1bec4630a
child 30466 5f31e24937c5
permissions -rw-r--r--
added dest_ss;
proper context fo pretty_ss;
tuned;
     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 meta_simplifier.ML for the actual meta-level rewriting engine).
     6 *)
     7 
     8 signature BASIC_SIMPLIFIER =
     9 sig
    10   include BASIC_META_SIMPLIFIER
    11   val change_simpset: (simpset -> simpset) -> unit
    12   val simpset_of: theory -> simpset
    13   val simpset: unit -> simpset
    14   val SIMPSET: (simpset -> tactic) -> tactic
    15   val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic
    16   val Addsimps: thm list -> unit
    17   val Delsimps: thm list -> unit
    18   val Addsimprocs: simproc list -> unit
    19   val Delsimprocs: simproc list -> unit
    20   val Addcongs: thm list -> unit
    21   val Delcongs: thm list -> unit
    22   val local_simpset_of: Proof.context -> simpset
    23   val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
    24   val safe_asm_full_simp_tac: simpset -> int -> tactic
    25   val               simp_tac: simpset -> int -> tactic
    26   val           asm_simp_tac: simpset -> int -> tactic
    27   val          full_simp_tac: simpset -> int -> tactic
    28   val        asm_lr_simp_tac: simpset -> int -> tactic
    29   val      asm_full_simp_tac: simpset -> int -> tactic
    30   val               Simp_tac:            int -> tactic
    31   val           Asm_simp_tac:            int -> tactic
    32   val          Full_simp_tac:            int -> tactic
    33   val        Asm_lr_simp_tac:            int -> tactic
    34   val      Asm_full_simp_tac:            int -> tactic
    35   val          simplify: simpset -> thm -> thm
    36   val      asm_simplify: simpset -> thm -> thm
    37   val     full_simplify: simpset -> thm -> thm
    38   val   asm_lr_simplify: simpset -> thm -> thm
    39   val asm_full_simplify: simpset -> thm -> thm
    40 end;
    41 
    42 signature SIMPLIFIER =
    43 sig
    44   include BASIC_SIMPLIFIER
    45   val pretty_ss: Proof.context -> simpset -> Pretty.T
    46   val clear_ss: simpset -> simpset
    47   val debug_bounds: bool ref
    48   val inherit_context: simpset -> simpset -> simpset
    49   val the_context: simpset -> Proof.context
    50   val context: Proof.context -> simpset -> simpset
    51   val theory_context: theory  -> simpset -> simpset
    52   val simproc_i: theory -> string -> term list
    53     -> (theory -> simpset -> term -> thm option) -> simproc
    54   val simproc: theory -> string -> string list
    55     -> (theory -> simpset -> term -> thm option) -> simproc
    56   val          rewrite: simpset -> conv
    57   val      asm_rewrite: simpset -> conv
    58   val     full_rewrite: simpset -> conv
    59   val   asm_lr_rewrite: simpset -> conv
    60   val asm_full_rewrite: simpset -> conv
    61   val get_ss: Context.generic -> simpset
    62   val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic
    63   val attrib: (simpset * thm list -> simpset) -> attribute
    64   val simp_add: attribute
    65   val simp_del: attribute
    66   val cong_add: attribute
    67   val cong_del: attribute
    68   val map_simpset: (simpset -> simpset) -> theory -> theory
    69   val get_simproc: Context.generic -> xstring -> simproc
    70   val def_simproc: {name: string, lhss: string list,
    71     proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    72     local_theory -> local_theory
    73   val def_simproc_i: {name: string, lhss: term list,
    74     proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->
    75     local_theory -> local_theory
    76   val cong_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
    77   val simp_modifiers': (Args.T list -> (Method.modifier * Args.T list)) list
    78   val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
    79   val method_setup: (Args.T list -> (Method.modifier * Args.T list)) list
    80     -> theory -> theory
    81   val easy_setup: thm -> thm list -> theory -> theory
    82 end;
    83 
    84 structure Simplifier: SIMPLIFIER =
    85 struct
    86 
    87 open MetaSimplifier;
    88 
    89 
    90 (** pretty printing **)
    91 
    92 fun pretty_ss ctxt ss =
    93   let
    94     val pretty_cterm = Syntax.pretty_term ctxt o Thm.term_of;
    95     val pretty_thm = ProofContext.pretty_thm ctxt;
    96     fun pretty_proc (name, lhss) = Pretty.big_list (name ^ ":") (map pretty_cterm lhss);
    97     fun pretty_cong (name, thm) =
    98       Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, pretty_thm thm];
    99 
   100     val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss;
   101   in
   102     [Pretty.big_list "simplification rules:" (map (pretty_thm o #2) simps),
   103       Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)),
   104       Pretty.big_list "congruences:" (map pretty_cong congs),
   105       Pretty.strs ("loopers:" :: map quote loopers),
   106       Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
   107       Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
   108     |> Pretty.chunks
   109   end;
   110 
   111 
   112 
   113 (** simpset data **)
   114 
   115 structure SimpsetData = GenericDataFun
   116 (
   117   type T = simpset;
   118   val empty = empty_ss;
   119   fun extend ss = MetaSimplifier.inherit_context empty_ss ss;
   120   fun merge _ = merge_ss;
   121 );
   122 
   123 val get_ss = SimpsetData.get;
   124 val map_ss = SimpsetData.map;
   125 
   126 
   127 (* attributes *)
   128 
   129 fun attrib f = Thm.declaration_attribute (fn th => map_ss (fn ss => f (ss, [th])));
   130 
   131 val simp_add = attrib (op addsimps);
   132 val simp_del = attrib (op delsimps);
   133 val cong_add = attrib (op addcongs);
   134 val cong_del = attrib (op delcongs);
   135 
   136 
   137 (* global simpset *)
   138 
   139 fun map_simpset f = Context.theory_map (map_ss f);
   140 fun change_simpset f = Context.>> (Context.map_theory (map_simpset f));
   141 fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_ss (Context.Theory thy));
   142 val simpset = simpset_of o ML_Context.the_global_context;
   143 
   144 fun SIMPSET tacf st = tacf (simpset_of (Thm.theory_of_thm st)) st;
   145 fun SIMPSET' tacf i st = tacf (simpset_of (Thm.theory_of_thm st)) i st;
   146 
   147 fun Addsimps args = change_simpset (fn ss => ss addsimps args);
   148 fun Delsimps args = change_simpset (fn ss => ss delsimps args);
   149 fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args);
   150 fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args);
   151 fun Addcongs args = change_simpset (fn ss => ss addcongs args);
   152 fun Delcongs args = change_simpset (fn ss => ss delcongs args);
   153 
   154 
   155 (* local simpset *)
   156 
   157 fun local_simpset_of ctxt = MetaSimplifier.context ctxt (get_ss (Context.Proof ctxt));
   158 
   159 val _ = ML_Antiquote.value "simpset"
   160   (Scan.succeed "Simplifier.local_simpset_of (ML_Context.the_local_context ())");
   161 
   162 
   163 
   164 (** named simprocs **)
   165 
   166 fun err_dup_simproc name = error ("Duplicate simproc: " ^ quote name);
   167 
   168 
   169 (* data *)
   170 
   171 structure Simprocs = GenericDataFun
   172 (
   173   type T = simproc NameSpace.table;
   174   val empty = NameSpace.empty_table;
   175   val extend = I;
   176   fun merge _ simprocs = NameSpace.merge_tables eq_simproc simprocs
   177     handle Symtab.DUP dup => err_dup_simproc dup;
   178 );
   179 
   180 
   181 (* get simprocs *)
   182 
   183 fun get_simproc context xname =
   184   let
   185     val (space, tab) = Simprocs.get context;
   186     val name = NameSpace.intern space xname;
   187   in
   188     (case Symtab.lookup tab name of
   189       SOME proc => proc
   190     | NONE => error ("Undefined simplification procedure: " ^ quote name))
   191   end;
   192 
   193 val _ = ML_Antiquote.value "simproc" (Scan.lift Args.name >> (fn name =>
   194   "Simplifier.get_simproc (ML_Context.the_generic_context ()) " ^ ML_Syntax.print_string name));
   195 
   196 
   197 (* define simprocs *)
   198 
   199 local
   200 
   201 fun gen_simproc prep {name, lhss, proc, identifier} lthy =
   202   let
   203     val b = Binding.name name;
   204     val naming = LocalTheory.full_naming lthy;
   205     val simproc = make_simproc
   206       {name = LocalTheory.full_name lthy b,
   207        lhss =
   208         let
   209           val lhss' = prep lthy lhss;
   210           val ctxt' = lthy
   211             |> fold Variable.declare_term lhss'
   212             |> fold Variable.auto_fixes lhss';
   213         in Variable.export_terms ctxt' lthy lhss' end
   214         |> map (Thm.cterm_of (ProofContext.theory_of lthy)),
   215        proc = proc,
   216        identifier = identifier}
   217       |> morph_simproc (LocalTheory.target_morphism lthy);
   218   in
   219     lthy |> LocalTheory.declaration (fn phi =>
   220       let
   221         val b' = Morphism.binding phi b;
   222         val simproc' = morph_simproc phi simproc;
   223       in
   224         Simprocs.map (fn simprocs =>
   225           NameSpace.bind naming (b', simproc') simprocs |> snd
   226             handle Symtab.DUP dup => err_dup_simproc dup)
   227         #> map_ss (fn ss => ss addsimprocs [simproc'])
   228       end)
   229   end;
   230 
   231 in
   232 
   233 val def_simproc = gen_simproc Syntax.read_terms;
   234 val def_simproc_i = gen_simproc Syntax.check_terms;
   235 
   236 end;
   237 
   238 
   239 
   240 (** simplification tactics and rules **)
   241 
   242 fun solve_all_tac solvers ss =
   243   let
   244     val (_, {subgoal_tac, ...}) = MetaSimplifier.internal_ss ss;
   245     val solve_tac = subgoal_tac (MetaSimplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac);
   246   in DEPTH_SOLVE (solve_tac 1) end;
   247 
   248 (*NOTE: may instantiate unknowns that appear also in other subgoals*)
   249 fun generic_simp_tac safe mode ss =
   250   let
   251     val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = MetaSimplifier.internal_ss ss;
   252     val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs));
   253     val solve_tac = FIRST' (map (MetaSimplifier.solver ss)
   254       (rev (if safe then solvers else unsafe_solvers)));
   255 
   256     fun simp_loop_tac i =
   257       asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN
   258       (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));
   259   in simp_loop_tac end;
   260 
   261 local
   262 
   263 fun simp rew mode ss thm =
   264   let
   265     val (_, {solvers = (unsafe_solvers, _), ...}) = MetaSimplifier.internal_ss ss;
   266     val tacf = solve_all_tac (rev unsafe_solvers);
   267     fun prover s th = Option.map #1 (Seq.pull (tacf s th));
   268   in rew mode prover ss thm end;
   269 
   270 in
   271 
   272 val simp_thm = simp MetaSimplifier.rewrite_thm;
   273 val simp_cterm = simp MetaSimplifier.rewrite_cterm;
   274 
   275 end;
   276 
   277 
   278 (* tactics *)
   279 
   280 val simp_tac = generic_simp_tac false (false, false, false);
   281 val asm_simp_tac = generic_simp_tac false (false, true, false);
   282 val full_simp_tac = generic_simp_tac false (true, false, false);
   283 val asm_lr_simp_tac = generic_simp_tac false (true, true, false);
   284 val asm_full_simp_tac = generic_simp_tac false (true, true, true);
   285 val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);
   286 
   287 (*the abstraction over the proof state delays the dereferencing*)
   288 fun          Simp_tac i st =          simp_tac (simpset ()) i st;
   289 fun      Asm_simp_tac i st =      asm_simp_tac (simpset ()) i st;
   290 fun     Full_simp_tac i st =     full_simp_tac (simpset ()) i st;
   291 fun   Asm_lr_simp_tac i st =   asm_lr_simp_tac (simpset ()) i st;
   292 fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st;
   293 
   294 
   295 (* conversions *)
   296 
   297 val          simplify = simp_thm (false, false, false);
   298 val      asm_simplify = simp_thm (false, true, false);
   299 val     full_simplify = simp_thm (true, false, false);
   300 val   asm_lr_simplify = simp_thm (true, true, false);
   301 val asm_full_simplify = simp_thm (true, true, true);
   302 
   303 val          rewrite = simp_cterm (false, false, false);
   304 val      asm_rewrite = simp_cterm (false, true, false);
   305 val     full_rewrite = simp_cterm (true, false, false);
   306 val   asm_lr_rewrite = simp_cterm (true, true, false);
   307 val asm_full_rewrite = simp_cterm (true, true, true);
   308 
   309 
   310 
   311 (** concrete syntax of attributes **)
   312 
   313 (* add / del *)
   314 
   315 val simpN = "simp";
   316 val congN = "cong";
   317 val addN = "add";
   318 val delN = "del";
   319 val onlyN = "only";
   320 val no_asmN = "no_asm";
   321 val no_asm_useN = "no_asm_use";
   322 val no_asm_simpN = "no_asm_simp";
   323 val asm_lrN = "asm_lr";
   324 
   325 
   326 (* simprocs *)
   327 
   328 local
   329 
   330 val add_del =
   331   (Args.del -- Args.colon >> K (op delsimprocs) ||
   332     Scan.option (Args.add -- Args.colon) >> K (op addsimprocs))
   333   >> (fn f => fn simproc => fn phi => Thm.declaration_attribute
   334       (K (map_ss (fn ss => f (ss, [morph_simproc phi simproc])))));
   335 
   336 in
   337 
   338 val simproc_att = Attrib.syntax
   339   (Scan.peek (fn context =>
   340     add_del :|-- (fn decl =>
   341       Scan.repeat1 (Args.named_attribute (decl o get_simproc context))
   342       >> (Library.apply o map Morphism.form))));
   343 
   344 end;
   345 
   346 
   347 (* conversions *)
   348 
   349 local
   350 
   351 fun conv_mode x =
   352   ((Args.parens (Args.$$$ no_asmN) >> K simplify ||
   353     Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify ||
   354     Args.parens (Args.$$$ no_asm_useN) >> K full_simplify ||
   355     Scan.succeed asm_full_simplify) |> Scan.lift) x;
   356 
   357 in
   358 
   359 val simplified =
   360   Attrib.syntax (conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute (fn context =>
   361     f ((if null ths then I else MetaSimplifier.clear_ss)
   362         (local_simpset_of (Context.proof_of context)) addsimps ths))));
   363 
   364 end;
   365 
   366 
   367 (* setup attributes *)
   368 
   369 val _ = Context.>> (Context.map_theory
   370  (Attrib.add_attributes
   371    [(simpN, Attrib.add_del_args simp_add simp_del, "declaration of Simplifier rewrite rule"),
   372     (congN, Attrib.add_del_args cong_add cong_del, "declaration of Simplifier congruence rule"),
   373     ("simproc", simproc_att, "declaration of simplification procedures"),
   374     ("simplified", simplified, "simplified rule")]));
   375 
   376 
   377 
   378 (** proof methods **)
   379 
   380 (* simplification *)
   381 
   382 val simp_options =
   383  (Args.parens (Args.$$$ no_asmN) >> K simp_tac ||
   384   Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac ||
   385   Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac ||
   386   Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac ||
   387   Scan.succeed asm_full_simp_tac);
   388 
   389 val cong_modifiers =
   390  [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),
   391   Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),
   392   Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];
   393 
   394 val simp_modifiers =
   395  [Args.$$$ simpN -- Args.colon >> K (I, simp_add),
   396   Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),
   397   Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),
   398   Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon
   399     >> K (Context.proof_map (map_ss MetaSimplifier.clear_ss), simp_add)]
   400    @ cong_modifiers;
   401 
   402 val simp_modifiers' =
   403  [Args.add -- Args.colon >> K (I, simp_add),
   404   Args.del -- Args.colon >> K (I, simp_del),
   405   Args.$$$ onlyN -- Args.colon
   406     >> K (Context.proof_map (map_ss MetaSimplifier.clear_ss), simp_add)]
   407    @ cong_modifiers;
   408 
   409 fun simp_args more_mods =
   410   Method.sectioned_args (Args.bang_facts -- Scan.lift simp_options)
   411     (more_mods @ simp_modifiers');
   412 
   413 fun simp_method (prems, tac) ctxt = Method.METHOD (fn facts =>
   414   ALLGOALS (Method.insert_tac (prems @ facts)) THEN
   415     (CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt));
   416 
   417 fun simp_method' (prems, tac) ctxt = Method.METHOD (fn facts =>
   418   HEADGOAL (Method.insert_tac (prems @ facts) THEN'
   419       ((CHANGED_PROP) oo tac) (local_simpset_of ctxt)));
   420 
   421 
   422 
   423 (** setup **)
   424 
   425 fun method_setup more_mods = Method.add_methods
   426  [(simpN, simp_args more_mods simp_method', "simplification"),
   427   ("simp_all", simp_args more_mods simp_method, "simplification (all goals)")];
   428 
   429 fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ =>
   430   let
   431     val trivialities = Drule.reflexive_thm :: trivs;
   432 
   433     fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac];
   434     val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;
   435 
   436     (*no premature instantiation of variables during simplification*)
   437     fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac];
   438     val safe_solver = mk_solver "easy safe" safe_solver_tac;
   439 
   440     fun mk_eq thm =
   441       if can Logic.dest_equals (Thm.concl_of thm) then [thm]
   442       else [thm RS reflect] handle THM _ => [];
   443 
   444     fun mksimps thm = mk_eq (Thm.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm);
   445   in
   446     empty_ss setsubgoaler asm_simp_tac
   447     setSSolver safe_solver
   448     setSolver unsafe_solver
   449     setmksimps mksimps
   450   end));
   451 
   452 end;
   453 
   454 structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier;
   455 open BasicSimplifier;