src/Provers/clasimp.ML
author wenzelm
Sun, 01 Mar 2009 23:36:12 +0100
changeset 30193 479806475f3c
parent 27809 a1e409db516b
child 30515 4120fc59dd85
permissions -rw-r--r--
use long names for old-style fold combinators;
     1 (*  Title:      Provers/clasimp.ML
     2     Author:     David von Oheimb, TU Muenchen
     3 
     4 Combination of classical reasoner and simplifier (depends on
     5 splitter.ML, classical.ML, blast.ML).
     6 *)
     7 
     8 infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 addsimps2 delsimps2
     9   addSss addss addss' addIffs delIffs;
    10 
    11 signature CLASIMP_DATA =
    12 sig
    13   structure Splitter: SPLITTER
    14   structure Classical: CLASSICAL
    15   structure Blast: BLAST
    16   sharing type Classical.claset = Blast.claset
    17   val notE: thm
    18   val iffD1: thm
    19   val iffD2: thm
    20 end
    21 
    22 signature CLASIMP =
    23 sig
    24   include CLASIMP_DATA
    25   type claset
    26   type clasimpset
    27   val get_css: Context.generic -> clasimpset
    28   val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic
    29   val change_clasimpset: (clasimpset -> clasimpset) -> unit
    30   val clasimpset: unit -> clasimpset
    31   val local_clasimpset_of: Proof.context -> clasimpset
    32   val addSIs2: clasimpset * thm list -> clasimpset
    33   val addSEs2: clasimpset * thm list -> clasimpset
    34   val addSDs2: clasimpset * thm list -> clasimpset
    35   val addIs2: clasimpset * thm list -> clasimpset
    36   val addEs2: clasimpset * thm list -> clasimpset
    37   val addDs2: clasimpset * thm list -> clasimpset
    38   val addsimps2: clasimpset * thm list -> clasimpset
    39   val delsimps2: clasimpset * thm list -> clasimpset
    40   val addSss: claset * simpset -> claset
    41   val addss: claset * simpset -> claset
    42   val addss': claset * simpset -> claset
    43   val addIffs: clasimpset * thm list -> clasimpset
    44   val delIffs: clasimpset * thm list -> clasimpset
    45   val AddIffs: thm list -> unit
    46   val DelIffs: thm list -> unit
    47   val CLASIMPSET: (clasimpset -> tactic) -> tactic
    48   val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic
    49   val clarsimp_tac: clasimpset -> int -> tactic
    50   val Clarsimp_tac: int -> tactic
    51   val mk_auto_tac: clasimpset -> int -> int -> tactic
    52   val auto_tac: clasimpset -> tactic
    53   val Auto_tac: tactic
    54   val auto: unit -> unit
    55   val force_tac: clasimpset  -> int -> tactic
    56   val Force_tac: int -> tactic
    57   val force: int -> unit
    58   val fast_simp_tac: clasimpset -> int -> tactic
    59   val slow_simp_tac: clasimpset -> int -> tactic
    60   val best_simp_tac: clasimpset -> int -> tactic
    61   val attrib: (clasimpset * thm list -> clasimpset) -> attribute
    62   val iff_add: attribute
    63   val iff_add': attribute
    64   val iff_del: attribute
    65   val iff_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
    66   val clasimp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
    67   val clasimp_setup: theory -> theory
    68 end;
    69 
    70 functor ClasimpFun(Data: CLASIMP_DATA): CLASIMP =
    71 struct
    72 
    73 open Data;
    74 
    75 
    76 (* type clasimpset *)
    77 
    78 type claset = Classical.claset;
    79 type clasimpset = claset * simpset;
    80 
    81 fun get_css context = (Classical.get_cs context, Simplifier.get_ss context);
    82 
    83 fun map_css f context =
    84   let val (cs', ss') = f (get_css context)
    85   in context |> Classical.map_cs (K cs') |> Simplifier.map_ss (K ss') end;
    86 
    87 fun change_clasimpset f = Context.>> (fn context => (Context.the_theory context; map_css f context));
    88 fun clasimpset () = (Classical.claset (), Simplifier.simpset ());
    89 
    90 fun local_clasimpset_of ctxt =
    91   (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt);
    92 
    93 
    94 (* clasimpset operations *)
    95 
    96 (*this interface for updating a clasimpset is rudimentary and just for
    97   convenience for the most common cases*)
    98 
    99 fun pair_upd1 f ((a, b), x) = (f (a, x), b);
   100 fun pair_upd2 f ((a, b), x) = (a, f (b, x));
   101 
   102 fun op addSIs2   arg = pair_upd1 Classical.addSIs arg;
   103 fun op addSEs2   arg = pair_upd1 Classical.addSEs arg;
   104 fun op addSDs2   arg = pair_upd1 Classical.addSDs arg;
   105 fun op addIs2    arg = pair_upd1 Classical.addIs arg;
   106 fun op addEs2    arg = pair_upd1 Classical.addEs arg;
   107 fun op addDs2    arg = pair_upd1 Classical.addDs arg;
   108 fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg;
   109 fun op delsimps2 arg = pair_upd2 Simplifier.delsimps arg;
   110 
   111 (*not totally safe: may instantiate unknowns that appear also in other subgoals*)
   112 val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true);
   113 
   114 (*Add a simpset to a classical set!*)
   115 (*Caution: only one simpset added can be added by each of addSss and addss*)
   116 fun cs addSss ss = Classical.addSafter (cs, ("safe_asm_full_simp_tac",
   117                             CHANGED o safe_asm_full_simp_tac ss));
   118 fun cs addss  ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
   119                             CHANGED o Simplifier.asm_full_simp_tac ss));
   120 fun cs addss' ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
   121                             CHANGED o Simplifier.asm_lr_simp_tac ss));
   122 
   123 (* iffs: addition of rules to simpsets and clasets simultaneously *)
   124 
   125 (*Takes (possibly conditional) theorems of the form A<->B to
   126         the Safe Intr     rule B==>A and
   127         the Safe Destruct rule A==>B.
   128   Also ~A goes to the Safe Elim rule A ==> ?R
   129   Failing other cases, A is added as a Safe Intr rule*)
   130 local
   131 
   132 fun addIff decls1 decls2 simp ((cs, ss), th) =
   133   let
   134     val n = nprems_of th;
   135     val (elim, intro) = if n = 0 then decls1 else decls2;
   136     val zero_rotate = zero_var_indexes o rotate_prems n;
   137   in
   138     (elim (intro (cs, [zero_rotate (th RS iffD2)]),
   139            [Tactic.make_elim (zero_rotate (th RS iffD1))])
   140       handle THM _ => (elim (cs, [zero_rotate (th RS notE)])
   141         handle THM _ => intro (cs, [th])),
   142      simp (ss, [th]))
   143   end;
   144 
   145 fun delIff delcs delss ((cs, ss), th) =
   146   let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in
   147     (delcs (cs, [zero_rotate (th RS iffD2),
   148         Tactic.make_elim (zero_rotate (th RS iffD1))])
   149       handle THM _ => (delcs (cs, [zero_rotate (th RS notE)])
   150         handle THM _ => delcs (cs, [th])),
   151      delss (ss, [th]))
   152   end;
   153 
   154 fun modifier att (x, ths) =
   155   fst (Library.foldl_map (Library.apply [att]) (x, rev ths));
   156 
   157 val addXIs = modifier (ContextRules.intro_query NONE);
   158 val addXEs = modifier (ContextRules.elim_query NONE);
   159 val addXDs = modifier (ContextRules.dest_query NONE);
   160 val delXs = modifier ContextRules.rule_del;
   161 
   162 in
   163 
   164 val op addIffs =
   165   Library.foldl
   166       (addIff (Classical.addSEs, Classical.addSIs)
   167               (Classical.addEs, Classical.addIs)
   168               Simplifier.addsimps);
   169 val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps);
   170 
   171 fun AddIffs thms = change_clasimpset (fn css => css addIffs thms);
   172 fun DelIffs thms = change_clasimpset (fn css => css delIffs thms);
   173 
   174 fun addIffs_generic (x, ths) =
   175   Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1;
   176 
   177 fun delIffs_generic (x, ths) =
   178   Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1;
   179 
   180 end;
   181 
   182 
   183 (* tacticals *)
   184 
   185 fun CLASIMPSET tacf state =
   186   Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state;
   187 
   188 fun CLASIMPSET' tacf i state =
   189   Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state;
   190 
   191 
   192 fun clarsimp_tac (cs, ss) =
   193   safe_asm_full_simp_tac ss THEN_ALL_NEW
   194   Classical.clarify_tac (cs addSss ss);
   195 
   196 fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i;
   197 
   198 
   199 (* auto_tac *)
   200 
   201 fun blast_depth_tac cs m i thm =
   202     Blast.depth_tac cs m i thm
   203       handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
   204 
   205 (* a variant of depth_tac that avoids interference of the simplifier
   206    with dup_step_tac when they are combined by auto_tac *)
   207 local
   208 fun slow_step_tac' cs = Classical.appWrappers cs
   209         (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs);
   210 in fun nodup_depth_tac cs m i state = SELECT_GOAL
   211    (Classical.safe_steps_tac cs 1 THEN_ELSE
   212         (DEPTH_SOLVE (nodup_depth_tac cs m 1),
   213          Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac
   214              (slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m-1) 1))
   215         )) i state;
   216 end;
   217 
   218 (*Designed to be idempotent, except if blast_depth_tac instantiates variables
   219   in some of the subgoals*)
   220 fun mk_auto_tac (cs, ss) m n =
   221     let val cs' = cs addss ss
   222         val maintac =
   223           blast_depth_tac cs m               (* fast but can't use wrappers *)
   224           ORELSE'
   225           (CHANGED o nodup_depth_tac cs' n); (* slower but more general *)
   226     in  EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss),
   227                TRY (Classical.safe_tac cs),
   228                REPEAT (FIRSTGOAL maintac),
   229                TRY (Classical.safe_tac (cs addSss ss)),
   230                prune_params_tac]
   231     end;
   232 
   233 fun auto_tac css = mk_auto_tac css 4 2;
   234 fun Auto_tac st = auto_tac (clasimpset ()) st;
   235 fun auto () = by Auto_tac;
   236 
   237 
   238 (* force_tac *)
   239 
   240 (* aimed to solve the given subgoal totally, using whatever tools possible *)
   241 fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [
   242         Classical.clarify_tac cs' 1,
   243         IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1),
   244         ALLGOALS (Classical.first_best_tac cs')]) end;
   245 fun Force_tac i = force_tac (clasimpset ()) i;
   246 fun force i = by (Force_tac i);
   247 
   248 
   249 (* basic combinations *)
   250 
   251 fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end;
   252 
   253 val fast_simp_tac = ADDSS Classical.fast_tac;
   254 val slow_simp_tac = ADDSS Classical.slow_tac;
   255 val best_simp_tac = ADDSS Classical.best_tac;
   256 
   257 
   258 (* access clasimpset *)
   259 
   260 
   261 
   262 (* attributes *)
   263 
   264 fun attrib f = Thm.declaration_attribute (fn th => map_css (fn css => f (css, [th])));
   265 fun attrib' f (x, th) = (f (x, [th]), th);
   266 
   267 val iff_add = attrib (op addIffs);
   268 val iff_add' = attrib' addIffs_generic;
   269 val iff_del = attrib (op delIffs) o attrib' delIffs_generic;
   270 
   271 val iff_att = Attrib.syntax (Scan.lift
   272  (Args.del >> K iff_del ||
   273   Scan.option Args.add -- Args.query >> K iff_add' ||
   274   Scan.option Args.add >> K iff_add));
   275 
   276 
   277 (* method modifiers *)
   278 
   279 val iffN = "iff";
   280 
   281 val iff_modifiers =
   282  [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier),
   283   Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'),
   284   Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)];
   285 
   286 val clasimp_modifiers =
   287   Simplifier.simp_modifiers @ Splitter.split_modifiers @
   288   Classical.cla_modifiers @ iff_modifiers;
   289 
   290 
   291 (* methods *)
   292 
   293 fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts =>
   294   ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt));
   295 
   296 fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts =>
   297   HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt)));
   298 
   299 val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth;
   300 val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth';
   301 
   302 
   303 fun auto_args m =
   304   Method.bang_sectioned_args' clasimp_modifiers
   305     (Scan.lift (Scan.option (OuterParse.nat -- OuterParse.nat))) m;
   306 
   307 fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac)
   308   | auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n));
   309 
   310 
   311 (* theory setup *)
   312 
   313 val clasimp_setup =
   314  (Attrib.add_attributes
   315    [("iff", iff_att, "declaration of Simplifier / Classical rules")] #>
   316   Method.add_methods
   317    [("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"),
   318     ("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"),
   319     ("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"),
   320     ("force", clasimp_method' force_tac, "force"),
   321     ("auto", auto_args auto_meth, "auto"),
   322     ("clarsimp", clasimp_method' (CHANGED_PROP oo clarsimp_tac), "clarify simplified goal")]);
   323 
   324 end;