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