src/Pure/Isar/find_theorems.ML
author wenzelm
Tue, 09 Oct 2007 00:20:13 +0200
changeset 24920 2a45e400fdad
parent 24683 c62320337a4e
child 25226 502d8676cdd6
permissions -rw-r--r--
generic Syntax.pretty/string_of operations;
     1 (*  Title:      Pure/Isar/find_theorems.ML
     2     ID:         $Id$
     3     Author:     Rafal Kolanski, NICTA and Tobias Nipkow, TU Muenchen
     4 
     5 Retrieve theorems from proof context.
     6 *)
     7 
     8 val thms_containing_limit = ref 40;
     9 
    10 signature FIND_THEOREMS =
    11 sig
    12   datatype 'term criterion =
    13     Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term
    14   val print_theorems: Proof.context -> term option -> int option -> bool ->
    15     (bool * string criterion) list -> unit
    16 end;
    17 
    18 structure FindTheorems: FIND_THEOREMS =
    19 struct
    20 
    21 (** search criteria **)
    22 
    23 datatype 'term criterion =
    24   Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term;
    25 
    26 fun read_criterion _ (Name name) = Name name
    27   | read_criterion _ Intro = Intro
    28   | read_criterion _ Elim = Elim
    29   | read_criterion _ Dest = Dest
    30   | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str)
    31   | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str);
    32 
    33 fun pretty_criterion ctxt (b, c) =
    34   let
    35     fun prfx s = if b then s else "-" ^ s;
    36   in
    37     (case c of
    38       Name name => Pretty.str (prfx "name: " ^ quote name)
    39     | Intro => Pretty.str (prfx "intro")
    40     | Elim => Pretty.str (prfx "elim")
    41     | Dest => Pretty.str (prfx "dest")
    42     | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,
    43         Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]
    44     | Pattern pat => Pretty.enclose (prfx " \"") "\""
    45         [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])
    46   end;
    47 
    48 
    49 
    50 (** search criterion filters **)
    51 
    52 (*generated filters are to be of the form
    53   input: (thmref * thm)
    54   output: (p:int, s:int) option, where
    55     NONE indicates no match
    56     p is the primary sorting criterion
    57       (eg. number of assumptions in the theorem)
    58     s is the secondary sorting criterion
    59       (eg. size of the substitution for intro, elim and dest)
    60   when applying a set of filters to a thm, fold results in:
    61     (biggest p, sum of all s)
    62   currently p and s only matter for intro, elim, dest and simp filters,
    63   otherwise the default ordering is used.
    64 *)
    65 
    66 
    67 (* matching theorems *)
    68 
    69 fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy;
    70 
    71 (*extract terms from term_src, refine them to the parts that concern us,
    72   if po try match them against obj else vice versa.
    73   trivial matches are ignored.
    74   returns: smallest substitution size*)
    75 fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =
    76   let
    77     val thy = ProofContext.theory_of ctxt;
    78 
    79     fun matches pat =
    80       is_nontrivial thy pat andalso
    81       Pattern.matches thy (if po then (pat, obj) else (obj, pat));
    82 
    83     fun substsize pat =
    84       let val (_, subst) =
    85         Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)
    86       in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
    87 
    88     fun bestmatch [] = NONE
    89      |  bestmatch xs = SOME (foldr1 Int.min xs);
    90 
    91     val match_thm = matches o refine_term;
    92   in
    93     map (substsize o refine_term)
    94         (filter match_thm (extract_terms term_src)) |> bestmatch
    95   end;
    96 
    97 
    98 (* filter_name *)
    99 
   100 fun match_string pat str =
   101   let
   102     fun match [] _ = true
   103       | match (p :: ps) s =
   104           size p <= size s andalso
   105             (case try (unprefix p) s of
   106               SOME s' => match ps s'
   107             | NONE => match (p :: ps) (String.substring (s, 1, size s - 1)));
   108   in match (space_explode "*" pat) str end;
   109 
   110 fun filter_name str_pat (thmref, _) =
   111   if match_string str_pat (PureThy.name_of_thmref thmref)
   112   then SOME (0, 0) else NONE;
   113 
   114 
   115 (* filter intro/elim/dest rules *)
   116 
   117 fun filter_dest ctxt goal (_, thm) =
   118   let
   119     val extract_dest =
   120      (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],
   121       hd o Logic.strip_imp_prems);
   122     val prems = Logic.prems_of_goal goal 1;
   123 
   124     fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm;
   125     val successful = prems |> map_filter try_subst;
   126   in
   127     (*if possible, keep best substitution (one with smallest size)*)
   128     (*dest rules always have assumptions, so a dest with one
   129       assumption is as good as an intro rule with none*)
   130     if not (null successful)
   131     then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
   132   end;
   133 
   134 fun filter_intro ctxt goal (_, thm) =
   135   let
   136     val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);
   137     val concl = Logic.concl_of_goal goal 1;
   138     val ss = is_matching_thm extract_intro ctxt true concl thm;
   139   in
   140     if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
   141   end;
   142 
   143 fun filter_elim ctxt goal (_, thm) =
   144   if not (Thm.no_prems thm) then
   145     let
   146       val rule = Thm.full_prop_of thm;
   147       val prems = Logic.prems_of_goal goal 1;
   148       val goal_concl = Logic.concl_of_goal goal 1;
   149       val rule_mp = (hd o Logic.strip_imp_prems) rule;
   150       val rule_concl = Logic.strip_imp_concl rule;
   151       fun combine t1 t2 = Const ("combine", dummyT --> dummyT) $ (t1 $ t2);
   152       val rule_tree = combine rule_mp rule_concl;
   153       fun goal_tree prem = (combine prem goal_concl);
   154       fun try_subst prem =
   155         is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree;
   156       val successful = prems |> map_filter try_subst;
   157     in
   158     (*elim rules always have assumptions, so an elim with one
   159       assumption is as good as an intro rule with none*)
   160       if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
   161         andalso not (null successful)
   162       then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
   163     end
   164   else NONE
   165 
   166 
   167 (* filter_simp *)
   168 
   169 fun filter_simp ctxt t (_, thm) =
   170   let
   171     val (_, {mk_rews = {mk, ...}, ...}) =
   172       MetaSimplifier.rep_ss (Simplifier.local_simpset_of ctxt);
   173     val extract_simp =
   174       (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl);
   175     val ss = is_matching_thm extract_simp ctxt false t thm
   176   in
   177     if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
   178   end;
   179 
   180 
   181 (* filter_pattern *)
   182 
   183 fun filter_pattern ctxt pat (_, thm) =
   184   if Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm)
   185   then SOME (0, 0) else NONE;
   186 
   187 
   188 (* interpret criteria as filters *)
   189 
   190 local
   191 
   192 fun err_no_goal c =
   193   error ("Current goal required for " ^ c ^ " search criterion");
   194 
   195 fun filter_crit _ _ (Name name) = filter_name name
   196   | filter_crit _ NONE Intro = err_no_goal "intro"
   197   | filter_crit _ NONE Elim = err_no_goal "elim"
   198   | filter_crit _ NONE Dest = err_no_goal "dest"
   199   | filter_crit ctxt (SOME goal) Intro = filter_intro ctxt goal
   200   | filter_crit ctxt (SOME goal) Elim = filter_elim ctxt goal
   201   | filter_crit ctxt (SOME goal) Dest = filter_dest ctxt goal
   202   | filter_crit ctxt _ (Simp pat) = filter_simp ctxt pat
   203   | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;
   204 
   205 fun opt_not x = if is_some x then NONE else SOME (0, 0);
   206 
   207 fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int)
   208  |  opt_add _ _ = NONE;
   209 
   210 in
   211 
   212 fun filter_criterion ctxt opt_goal (b, c) =
   213   (if b then I else opt_not) o filter_crit ctxt opt_goal c;
   214 
   215 fun all_filters filters thms =
   216   let
   217     fun eval_filters filters thm =
   218       fold opt_add (map (fn f => f thm) filters) (SOME (0, 0));
   219 
   220     (*filters return: (number of assumptions, substitution size) option, so
   221       sort (desc. in both cases) according to number of assumptions first,
   222       then by the substitution size*)
   223     fun thm_ord (((p0, s0), _), ((p1, s1), _)) =
   224       prod_ord int_ord int_ord ((p1, s1), (p0, s0));
   225   in
   226     map (`(eval_filters filters)) thms
   227     |> map_filter (fn (SOME x, y) => SOME (x, y) | (NONE, _) => NONE)
   228     |> sort thm_ord |> map #2
   229   end;
   230 
   231 end;
   232 
   233 
   234 (* removing duplicates, preferring nicer names, roughly n log n *)
   235 
   236 fun rem_thm_dups xs =
   237   let
   238     fun nicer (Fact x) (Fact y) = size x <= size y
   239       | nicer (Fact _) _        = true
   240       | nicer (PureThy.Name x) (PureThy.Name y) = size x <= size y
   241       | nicer (PureThy.Name _) (Fact _) = false
   242       | nicer (PureThy.Name _) _ = true
   243       | nicer (NameSelection (x, _)) (NameSelection (y, _)) = size x <= size y
   244       | nicer (NameSelection _) _ = false;
   245 
   246     fun rem_cdups xs =
   247       let
   248         fun rem_c rev_seen [] = rev rev_seen
   249           | rem_c rev_seen [x] = rem_c (x::rev_seen) []
   250           | rem_c rev_seen ((x as ((n,t),_))::(y as ((n',t'),_))::xs) =
   251             if Thm.eq_thm_prop (t,t')
   252             then if nicer n n'
   253                  then rem_c rev_seen (x::xs)
   254                  else rem_c rev_seen (y::xs)
   255             else rem_c (x::rev_seen) (y::xs)
   256       in rem_c [] xs end;
   257 
   258   in ListPair.zip (xs, 1 upto length xs)
   259      |> sort (Term.fast_term_ord o pairself (prop_of o #2 o #1))
   260      |> rem_cdups
   261      |> sort (int_ord o pairself #2)
   262      |> map #1
   263   end;
   264 
   265 
   266 (* print_theorems *)
   267 
   268 fun find_thms ctxt spec =
   269   (PureThy.thms_containing (ProofContext.theory_of ctxt) spec
   270     |> maps PureThy.selections) @
   271   (ProofContext.lthms_containing ctxt spec
   272     |> maps PureThy.selections
   273     |> distinct (fn ((r1, th1), (r2, th2)) =>
   274         r1 = r2 andalso Thm.eq_thm_prop (th1, th2)));
   275 
   276 fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria =
   277   let
   278     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
   279     val filters = map (filter_criterion ctxt opt_goal) criteria;
   280 
   281     val raw_matches = all_filters filters (find_thms ctxt ([], []));
   282     val matches =
   283       if rem_dups
   284       then rem_thm_dups raw_matches
   285       else raw_matches;
   286 
   287     val len = length matches;
   288     val limit = the_default (! thms_containing_limit) opt_limit;
   289     val thms = Library.drop (len - limit, matches);
   290 
   291     fun prt_fact (thmref, thm) =
   292       ProofContext.pretty_fact ctxt (PureThy.string_of_thmref thmref, [thm]);
   293   in
   294     Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) :: Pretty.str "" ::
   295      (if null thms then [Pretty.str "nothing found"]
   296       else
   297         [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^
   298           (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":"),
   299          Pretty.str ""] @
   300         map prt_fact thms)
   301     |> Pretty.chunks |> Pretty.writeln
   302   end;
   303 
   304 end;