src/Pure/Tools/find_theorems.ML
author wenzelm
Sat, 15 May 2010 23:40:00 +0200
changeset 36953 2af1ad9aa1a3
parent 36950 75b8f26f2f07
child 38593 c677c2c1d333
permissions -rw-r--r--
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
     1 (*  Title:      Pure/Tools/find_theorems.ML
     2     Author:     Rafal Kolanski and Gerwin Klein, NICTA
     3 
     4 Retrieve theorems from proof context.
     5 *)
     6 
     7 signature FIND_THEOREMS =
     8 sig
     9   datatype 'term criterion =
    10     Name of string | Intro | IntroIff | Elim | Dest | Solves | Simp of 'term |
    11     Pattern of 'term
    12   val tac_limit: int Unsynchronized.ref
    13   val limit: int Unsynchronized.ref
    14   val find_theorems: Proof.context -> thm option -> int option -> bool ->
    15     (bool * string criterion) list -> int option * (Facts.ref * thm) list
    16   val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T
    17   val print_theorems: Proof.context -> thm option -> int option -> bool ->
    18     (bool * string criterion) list -> unit
    19 end;
    20 
    21 structure Find_Theorems: FIND_THEOREMS =
    22 struct
    23 
    24 (** search criteria **)
    25 
    26 datatype 'term criterion =
    27   Name of string | Intro | IntroIff | Elim | Dest | Solves | Simp of 'term |
    28   Pattern of 'term;
    29 
    30 fun apply_dummies tm =
    31   let
    32     val (xs, _) = Term.strip_abs tm;
    33     val tm' = Term.betapplys (tm, map (Term.dummy_pattern o #2) xs);
    34   in #1 (Term.replace_dummy_patterns tm' 1) end;
    35 
    36 fun parse_pattern ctxt nm =
    37   let
    38     val consts = ProofContext.consts_of ctxt;
    39     val nm' =
    40       (case Syntax.parse_term ctxt nm of
    41         Const (c, _) => c
    42       | _ => Consts.intern consts nm);
    43   in
    44     (case try (Consts.the_abbreviation consts) nm' of
    45       SOME (_, rhs) => apply_dummies (ProofContext.expand_abbrevs ctxt rhs)
    46     | NONE => ProofContext.read_term_pattern ctxt nm)
    47   end;
    48 
    49 fun read_criterion _ (Name name) = Name name
    50   | read_criterion _ Intro = Intro
    51   | read_criterion _ IntroIff = IntroIff
    52   | read_criterion _ Elim = Elim
    53   | read_criterion _ Dest = Dest
    54   | read_criterion _ Solves = Solves
    55   | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str)
    56   | read_criterion ctxt (Pattern str) = Pattern (parse_pattern ctxt str);
    57 
    58 fun pretty_criterion ctxt (b, c) =
    59   let
    60     fun prfx s = if b then s else "-" ^ s;
    61   in
    62     (case c of
    63       Name name => Pretty.str (prfx "name: " ^ quote name)
    64     | Intro => Pretty.str (prfx "intro")
    65     | IntroIff => Pretty.str (prfx "introiff")
    66     | Elim => Pretty.str (prfx "elim")
    67     | Dest => Pretty.str (prfx "dest")
    68     | Solves => Pretty.str (prfx "solves")
    69     | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,
    70         Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]
    71     | Pattern pat => Pretty.enclose (prfx " \"") "\""
    72         [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])
    73   end;
    74 
    75 
    76 
    77 (** search criterion filters **)
    78 
    79 (*generated filters are to be of the form
    80   input: (Facts.ref * thm)
    81   output: (p:int, s:int) option, where
    82     NONE indicates no match
    83     p is the primary sorting criterion
    84       (eg. number of assumptions in the theorem)
    85     s is the secondary sorting criterion
    86       (eg. size of the substitution for intro, elim and dest)
    87   when applying a set of filters to a thm, fold results in:
    88     (biggest p, sum of all s)
    89   currently p and s only matter for intro, elim, dest and simp filters,
    90   otherwise the default ordering is used.
    91 *)
    92 
    93 
    94 (* matching theorems *)
    95 
    96 fun is_nontrivial thy = Term.is_Const o Term.head_of o Object_Logic.drop_judgment thy;
    97 
    98 (*educated guesses on HOL*)  (* FIXME broken *)
    99 val boolT = Type ("bool", []);
   100 val iff_const = Const ("op =", boolT --> boolT --> boolT);
   101 
   102 (*extract terms from term_src, refine them to the parts that concern us,
   103   if po try match them against obj else vice versa.
   104   trivial matches are ignored.
   105   returns: smallest substitution size*)
   106 fun is_matching_thm doiff (extract_terms, refine_term) ctxt po obj term_src =
   107   let
   108     val thy = ProofContext.theory_of ctxt;
   109 
   110     fun check_match pat = Pattern.matches thy (if po then (pat, obj) else (obj, pat));
   111     fun matches pat =
   112       let
   113         val jpat = Object_Logic.drop_judgment thy pat;
   114         val c = Term.head_of jpat;
   115         val pats =
   116           if Term.is_Const c
   117           then
   118             if doiff andalso c = iff_const then
   119               (pat :: map (Object_Logic.ensure_propT thy) (snd (strip_comb jpat)))
   120                 |> filter (is_nontrivial thy)
   121             else [pat]
   122           else [];
   123       in filter check_match pats end;
   124 
   125     fun substsize pat =
   126       let val (_, subst) =
   127         Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)
   128       in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
   129 
   130     fun bestmatch [] = NONE
   131       | bestmatch xs = SOME (foldl1 Int.min xs);
   132 
   133     val match_thm = matches o refine_term;
   134   in
   135     maps match_thm (extract_terms term_src)
   136     |> map substsize
   137     |> bestmatch
   138   end;
   139 
   140 
   141 (* filter_name *)
   142 
   143 fun filter_name str_pat (thmref, _) =
   144   if match_string str_pat (Facts.name_of_ref thmref)
   145   then SOME (0, 0) else NONE;
   146 
   147 
   148 (* filter intro/elim/dest/solves rules *)
   149 
   150 fun filter_dest ctxt goal (_, thm) =
   151   let
   152     val extract_dest =
   153      (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],
   154       hd o Logic.strip_imp_prems);
   155     val prems = Logic.prems_of_goal goal 1;
   156 
   157     fun try_subst prem = is_matching_thm false extract_dest ctxt true prem thm;
   158     val successful = prems |> map_filter try_subst;
   159   in
   160     (*if possible, keep best substitution (one with smallest size)*)
   161     (*dest rules always have assumptions, so a dest with one
   162       assumption is as good as an intro rule with none*)
   163     if not (null successful)
   164     then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE
   165   end;
   166 
   167 fun filter_intro doiff ctxt goal (_, thm) =
   168   let
   169     val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);
   170     val concl = Logic.concl_of_goal goal 1;
   171     val ss = is_matching_thm doiff extract_intro ctxt true concl thm;
   172   in
   173     if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
   174   end;
   175 
   176 fun filter_elim ctxt goal (_, thm) =
   177   if not (Thm.no_prems thm) then
   178     let
   179       val rule = Thm.full_prop_of thm;
   180       val prems = Logic.prems_of_goal goal 1;
   181       val goal_concl = Logic.concl_of_goal goal 1;
   182       val rule_mp = hd (Logic.strip_imp_prems rule);
   183       val rule_concl = Logic.strip_imp_concl rule;
   184       fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);
   185       val rule_tree = combine rule_mp rule_concl;
   186       fun goal_tree prem = combine prem goal_concl;
   187       fun try_subst prem =
   188         is_matching_thm false (single, I) ctxt true (goal_tree prem) rule_tree;
   189       val successful = prems |> map_filter try_subst;
   190     in
   191       (*elim rules always have assumptions, so an elim with one
   192         assumption is as good as an intro rule with none*)
   193       if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
   194         andalso not (null successful)
   195       then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE
   196     end
   197   else NONE
   198 
   199 val tac_limit = Unsynchronized.ref 5;
   200 
   201 fun filter_solves ctxt goal =
   202   let
   203     fun etacn thm i = Seq.take (! tac_limit) o etac thm i;
   204     fun try_thm thm =
   205       if Thm.no_prems thm then rtac thm 1 goal
   206       else (etacn thm THEN_ALL_NEW (Goal.norm_hhf_tac THEN' Method.assm_tac ctxt)) 1 goal;
   207   in
   208     fn (_, thm) =>
   209       if is_some (Seq.pull (try_thm thm))
   210       then SOME (Thm.nprems_of thm, 0) else NONE
   211   end;
   212 
   213 
   214 (* filter_simp *)
   215 
   216 fun filter_simp ctxt t (_, thm) =
   217   let
   218     val mksimps = Simplifier.mksimps (simpset_of ctxt);
   219     val extract_simp =
   220       (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl);
   221     val ss = is_matching_thm false extract_simp ctxt false t thm;
   222   in
   223     if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
   224   end;
   225 
   226 
   227 (* filter_pattern *)
   228 
   229 fun get_names t = Term.add_const_names t (Term.add_free_names t []);
   230 fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm);
   231 
   232 (*Including all constants and frees is only sound because
   233   matching uses higher-order patterns. If full matching
   234   were used, then constants that may be subject to
   235   beta-reduction after substitution of frees should
   236   not be included for LHS set because they could be
   237   thrown away by the substituted function.
   238   e.g. for (?F 1 2) do not include 1 or 2, if it were
   239        possible for ?F to be (% x y. 3)
   240   The largest possible set should always be included on
   241   the RHS.*)
   242 
   243 fun filter_pattern ctxt pat =
   244   let
   245     val pat_consts = get_names pat;
   246 
   247     fun check (t, NONE) = check (t, SOME (get_thm_names t))
   248       | check ((_, thm), c as SOME thm_consts) =
   249          (if subset (op =) (pat_consts, thm_consts) andalso
   250             Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm)
   251           then SOME (0, 0) else NONE, c);
   252   in check end;
   253 
   254 
   255 (* interpret criteria as filters *)
   256 
   257 local
   258 
   259 fun err_no_goal c =
   260   error ("Current goal required for " ^ c ^ " search criterion");
   261 
   262 val fix_goal = Thm.prop_of;
   263 
   264 fun filter_crit _ _ (Name name) = apfst (filter_name name)
   265   | filter_crit _ NONE Intro = err_no_goal "intro"
   266   | filter_crit _ NONE Elim = err_no_goal "elim"
   267   | filter_crit _ NONE Dest = err_no_goal "dest"
   268   | filter_crit _ NONE Solves = err_no_goal "solves"
   269   | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro false ctxt (fix_goal goal))
   270   | filter_crit ctxt (SOME goal) IntroIff = apfst (filter_intro true ctxt (fix_goal goal))
   271   | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (fix_goal goal))
   272   | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (fix_goal goal))
   273   | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal)
   274   | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat)
   275   | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;
   276 
   277 fun opt_not x = if is_some x then NONE else SOME (0, 0);
   278 
   279 fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int)
   280   | opt_add _ _ = NONE;
   281 
   282 fun app_filters thm =
   283   let
   284     fun app (NONE, _, _) = NONE
   285       | app (SOME v, _, []) = SOME (v, thm)
   286       | app (r, consts, f :: fs) =
   287           let val (r', consts') = f (thm, consts)
   288           in app (opt_add r r', consts', fs) end;
   289   in app end;
   290 
   291 
   292 in
   293 
   294 fun filter_criterion ctxt opt_goal (b, c) =
   295   (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c;
   296 
   297 fun sorted_filter filters thms =
   298   let
   299     fun eval_filters thm = app_filters thm (SOME (0, 0), NONE, filters);
   300 
   301     (*filters return: (number of assumptions, substitution size) option, so
   302       sort (desc. in both cases) according to number of assumptions first,
   303       then by the substitution size*)
   304     fun thm_ord (((p0, s0), _), ((p1, s1), _)) =
   305       prod_ord int_ord int_ord ((p1, s1), (p0, s0));
   306   in map_filter eval_filters thms |> sort thm_ord |> map #2 end;
   307 
   308 fun lazy_filter filters =
   309   let
   310     fun lazy_match thms = Seq.make (fn () => first_match thms)
   311 
   312     and first_match [] = NONE
   313       | first_match (thm :: thms) =
   314           (case app_filters thm (SOME (0, 0), NONE, filters) of
   315             NONE => first_match thms
   316           | SOME (_, t) => SOME (t, lazy_match thms));
   317   in lazy_match end;
   318 
   319 end;
   320 
   321 
   322 (* removing duplicates, preferring nicer names, roughly n log n *)
   323 
   324 local
   325 
   326 val index_ord = option_ord (K EQUAL);
   327 val hidden_ord = bool_ord o pairself Name_Space.is_hidden;
   328 val qual_ord = int_ord o pairself (length o Long_Name.explode);
   329 val txt_ord = int_ord o pairself size;
   330 
   331 fun nicer_name (x, i) (y, j) =
   332   (case hidden_ord (x, y) of EQUAL =>
   333     (case index_ord (i, j) of EQUAL =>
   334       (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord)
   335     | ord => ord)
   336   | ord => ord) <> GREATER;
   337 
   338 fun rem_cdups nicer xs =
   339   let
   340     fun rem_c rev_seen [] = rev rev_seen
   341       | rem_c rev_seen [x] = rem_c (x :: rev_seen) []
   342       | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) =
   343           if Thm.eq_thm_prop (t, t')
   344           then rem_c rev_seen ((if nicer n n' then x else y) :: xs)
   345           else rem_c (x :: rev_seen) (y :: xs)
   346   in rem_c [] xs end;
   347 
   348 in
   349 
   350 fun nicer_shortest ctxt =
   351   let
   352     (* FIXME global name space!? *)
   353     val space = Facts.space_of (PureThy.facts_of (ProofContext.theory_of ctxt));
   354 
   355     val shorten =
   356       Name_Space.extern_flags
   357         {long_names = false, short_names = false, unique_names = false} space;
   358 
   359     fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =
   360           nicer_name (shorten x, i) (shorten y, j)
   361       | nicer (Facts.Fact _) (Facts.Named _) = true
   362       | nicer (Facts.Named _) (Facts.Fact _) = false;
   363   in nicer end;
   364 
   365 fun rem_thm_dups nicer xs =
   366   xs ~~ (1 upto length xs)
   367   |> sort (Term_Ord.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
   368   |> rem_cdups nicer
   369   |> sort (int_ord o pairself #2)
   370   |> map #1;
   371 
   372 end;
   373 
   374 
   375 (* print_theorems *)
   376 
   377 fun all_facts_of ctxt =
   378   let
   379     fun visible_facts facts =
   380       Facts.dest_static [] facts
   381       |> filter_out (Facts.is_concealed facts o #1);
   382   in
   383     maps Facts.selections
   384      (visible_facts (PureThy.facts_of (ProofContext.theory_of ctxt)) @
   385       visible_facts (ProofContext.facts_of ctxt))
   386   end;
   387 
   388 val limit = Unsynchronized.ref 40;
   389 
   390 fun find_theorems ctxt opt_goal opt_limit rem_dups raw_criteria =
   391   let
   392     val assms =
   393       ProofContext.get_fact ctxt (Facts.named "local.assms")
   394         handle ERROR _ => [];
   395     val add_prems = Seq.hd o TRY (Method.insert_tac assms 1);
   396     val opt_goal' = Option.map add_prems opt_goal;
   397 
   398     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
   399     val filters = map (filter_criterion ctxt opt_goal') criteria;
   400 
   401     fun find_all facts =
   402       let
   403         val raw_matches = sorted_filter filters facts;
   404 
   405         val matches =
   406           if rem_dups
   407           then rem_thm_dups (nicer_shortest ctxt) raw_matches
   408           else raw_matches;
   409 
   410         val len = length matches;
   411         val lim = the_default (! limit) opt_limit;
   412       in (SOME len, drop (Int.max (len - lim, 0)) matches) end;
   413 
   414     val find =
   415       if rem_dups orelse is_none opt_limit
   416       then find_all
   417       else pair NONE o Seq.list_of o Seq.take (the opt_limit) o lazy_filter filters;
   418 
   419   in find (all_facts_of ctxt) end;
   420 
   421 
   422 fun pretty_thm ctxt (thmref, thm) = Pretty.block
   423   [Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1,
   424     Display.pretty_thm ctxt thm];
   425 
   426 fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria =
   427   let
   428     val start = start_timing ();
   429 
   430     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
   431     val (foundo, thms) = find_theorems ctxt opt_goal opt_limit rem_dups raw_criteria;
   432     val returned = length thms;
   433 
   434     val tally_msg =
   435       (case foundo of
   436         NONE => "displaying " ^ string_of_int returned ^ " theorems"
   437       | SOME found =>
   438           "found " ^ string_of_int found ^ " theorems" ^
   439             (if returned < found
   440              then " (" ^ string_of_int returned ^ " displayed)"
   441              else ""));
   442 
   443     val end_msg = " in " ^ Time.toString (#cpu (end_timing start)) ^ " secs";
   444   in
   445     Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria)
   446         :: Pretty.str "" ::
   447      (if null thms then [Pretty.str ("nothing found" ^ end_msg)]
   448       else
   449         [Pretty.str (tally_msg ^ end_msg ^ ":"), Pretty.str ""] @
   450         map (pretty_thm ctxt) thms)
   451     |> Pretty.chunks |> Pretty.writeln
   452   end;
   453 
   454 
   455 
   456 (** command syntax **)
   457 
   458 fun find_theorems_cmd ((opt_lim, rem_dups), spec) =
   459   Toplevel.unknown_theory o Toplevel.keep (fn state =>
   460     let
   461       val proof_state = Toplevel.enter_proof_body state;
   462       val ctxt = Proof.context_of proof_state;
   463       val opt_goal = try Proof.simple_goal proof_state |> Option.map #goal;
   464     in print_theorems ctxt opt_goal opt_lim rem_dups spec end);
   465 
   466 local
   467 
   468 val criterion =
   469   Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name ||
   470   Parse.reserved "intro" >> K Intro ||
   471   Parse.reserved "introiff" >> K IntroIff ||
   472   Parse.reserved "elim" >> K Elim ||
   473   Parse.reserved "dest" >> K Dest ||
   474   Parse.reserved "solves" >> K Solves ||
   475   Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp ||
   476   Parse.term >> Pattern;
   477 
   478 val options =
   479   Scan.optional
   480     (Parse.$$$ "(" |--
   481       Parse.!!! (Scan.option Parse.nat -- Scan.optional (Parse.reserved "with_dups" >> K false) true
   482         --| Parse.$$$ ")")) (NONE, true);
   483 in
   484 
   485 val _ =
   486   Outer_Syntax.improper_command "find_theorems" "print theorems meeting specified criteria"
   487     Keyword.diag
   488     (options -- Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion))
   489       >> (Toplevel.no_timing oo find_theorems_cmd));
   490 
   491 end;
   492 
   493 end;