(* Title: Pure/Isar/find_theorems.ML ID: $Id$ Author: Rafal Kolanski and Gerwin Klein, NICTARetrieve theorems from proof context.*)signature FIND_THEOREMS =sig val limit: int ref datatype 'term criterion = Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term val print_theorems: Proof.context -> term option -> int option -> bool -> (bool * string criterion) list -> unitend;structure FindTheorems: FIND_THEOREMS =struct(** search criteria **)datatype 'term criterion = Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term;fun read_criterion _ (Name name) = Name name | read_criterion _ Intro = Intro | read_criterion _ Elim = Elim | read_criterion _ Dest = Dest | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str) | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str);fun pretty_criterion ctxt (b, c) = let fun prfx s = if b then s else "-" ^ s; in (case c of Name name => Pretty.str (prfx "name: " ^ quote name) | Intro => Pretty.str (prfx "intro") | Elim => Pretty.str (prfx "elim") | Dest => Pretty.str (prfx "dest") | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] | Pattern pat => Pretty.enclose (prfx " \"") "\"" [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) end;(** search criterion filters **)(*generated filters are to be of the form input: (Facts.ref * thm) output: (p:int, s:int) option, where NONE indicates no match p is the primary sorting criterion (eg. number of assumptions in the theorem) s is the secondary sorting criterion (eg. size of the substitution for intro, elim and dest) when applying a set of filters to a thm, fold results in: (biggest p, sum of all s) currently p and s only matter for intro, elim, dest and simp filters, otherwise the default ordering is used.*)(* matching theorems *)fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy;(*extract terms from term_src, refine them to the parts that concern us, if po try match them against obj else vice versa. trivial matches are ignored. returns: smallest substitution size*)fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src = let val thy = ProofContext.theory_of ctxt; fun matches pat = is_nontrivial thy pat andalso Pattern.matches thy (if po then (pat, obj) else (obj, pat)); fun substsize pat = let val (_, subst) = Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty) in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end; fun bestmatch [] = NONE | bestmatch xs = SOME (foldr1 Int.min xs); val match_thm = matches o refine_term; in map (substsize o refine_term) (filter match_thm (extract_terms term_src)) |> bestmatch end;(* filter_name *)fun match_string pat str = let fun match [] _ = true | match (p :: ps) s = size p <= size s andalso (case try (unprefix p) s of SOME s' => match ps s' | NONE => match (p :: ps) (String.substring (s, 1, size s - 1))); in match (space_explode "*" pat) str end;fun filter_name str_pat (thmref, _) = if match_string str_pat (Facts.name_of_ref thmref) then SOME (0, 0) else NONE;(* filter intro/elim/dest rules *)fun filter_dest ctxt goal (_, thm) = let val extract_dest = (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm], hd o Logic.strip_imp_prems); val prems = Logic.prems_of_goal goal 1; fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm; val successful = prems |> map_filter try_subst; in (*if possible, keep best substitution (one with smallest size)*) (*dest rules always have assumptions, so a dest with one assumption is as good as an intro rule with none*) if not (null successful) then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE end;fun filter_intro ctxt goal (_, thm) = let val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl); val concl = Logic.concl_of_goal goal 1; val ss = is_matching_thm extract_intro ctxt true concl thm; in if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE end;fun filter_elim ctxt goal (_, thm) = if not (Thm.no_prems thm) then let val rule = Thm.full_prop_of thm; val prems = Logic.prems_of_goal goal 1; val goal_concl = Logic.concl_of_goal goal 1; val rule_mp = hd (Logic.strip_imp_prems rule); val rule_concl = Logic.strip_imp_concl rule; fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2); val rule_tree = combine rule_mp rule_concl; fun goal_tree prem = combine prem goal_concl; fun try_subst prem = is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree; val successful = prems |> map_filter try_subst; in (*elim rules always have assumptions, so an elim with one assumption is as good as an intro rule with none*) if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm) andalso not (null successful) then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE end else NONE(* filter_simp *)fun filter_simp ctxt t (_, thm) = let val (_, {mk_rews = {mk, ...}, ...}) = MetaSimplifier.rep_ss (Simplifier.local_simpset_of ctxt); val extract_simp = (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl); val ss = is_matching_thm extract_simp ctxt false t thm in if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE end;(* filter_pattern *)fun filter_pattern ctxt pat (_, thm) = if Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm) then SOME (0, 0) else NONE;(* interpret criteria as filters *)localfun err_no_goal c = error ("Current goal required for " ^ c ^ " search criterion");fun filter_crit _ _ (Name name) = filter_name name | filter_crit _ NONE Intro = err_no_goal "intro" | filter_crit _ NONE Elim = err_no_goal "elim" | filter_crit _ NONE Dest = err_no_goal "dest" | filter_crit ctxt (SOME goal) Intro = filter_intro ctxt goal | filter_crit ctxt (SOME goal) Elim = filter_elim ctxt goal | filter_crit ctxt (SOME goal) Dest = filter_dest ctxt goal | filter_crit ctxt _ (Simp pat) = filter_simp ctxt pat | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;fun opt_not x = if is_some x then NONE else SOME (0, 0);fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int) | opt_add _ _ = NONE;infun filter_criterion ctxt opt_goal (b, c) = (if b then I else opt_not) o filter_crit ctxt opt_goal c;fun all_filters filters thms = let fun eval_filters thm = fold opt_add (map (fn f => f thm) filters) (SOME (0, 0)); (*filters return: (number of assumptions, substitution size) option, so sort (desc. in both cases) according to number of assumptions first, then by the substitution size*) fun thm_ord (((p0, s0), _), ((p1, s1), _)) = prod_ord int_ord int_ord ((p1, s1), (p0, s0)); in map (`eval_filters) thms |> map_filter (fn (SOME x, y) => SOME (x, y) | (NONE, _) => NONE) |> sort thm_ord |> map #2 end;end;(* removing duplicates, preferring nicer names, roughly n log n *)localval index_ord = option_ord (K EQUAL);val hidden_ord = bool_ord o pairself NameSpace.is_hidden;val qual_ord = int_ord o pairself (length o NameSpace.explode);val txt_ord = int_ord o pairself size;fun nicer_name (x, i) (y, j) = (case hidden_ord (x, y) of EQUAL => (case index_ord (i, j) of EQUAL => (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord) | ord => ord) | ord => ord) <> GREATER;fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) = nicer_name (x, i) (y, j) | nicer (Facts.Fact _) (Facts.Named _) = true | nicer (Facts.Named _) (Facts.Fact _) = false;fun rem_cdups xs = let fun rem_c rev_seen [] = rev rev_seen | rem_c rev_seen [x] = rem_c (x :: rev_seen) [] | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) = if Thm.eq_thm_prop (t, t') then rem_c rev_seen ((if nicer n n' then x else y) :: xs) else rem_c (x :: rev_seen) (y :: xs) in rem_c [] xs end;infun rem_thm_dups xs = xs ~~ (1 upto length xs) |> sort (Term.fast_term_ord o pairself (Thm.prop_of o #2 o #1)) |> rem_cdups |> sort (int_ord o pairself #2) |> map #1;end;(* print_theorems *)fun all_facts_of ctxt = maps Facts.selections (Facts.dest_static [] (PureThy.facts_of (ProofContext.theory_of ctxt)) @ Facts.dest_static [] (ProofContext.facts_of ctxt));val limit = ref 40;fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = let val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; val filters = map (filter_criterion ctxt opt_goal) criteria; val raw_matches = all_filters filters (all_facts_of ctxt); val matches = if rem_dups then rem_thm_dups raw_matches else raw_matches; val len = length matches; val lim = the_default (! limit) opt_limit; val thms = Library.drop (len - lim, matches); fun prt_fact (thmref, thm) = ProofContext.pretty_fact ctxt (Facts.string_of_ref thmref, [thm]); in Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) :: Pretty.str "" :: (if null thms then [Pretty.str "nothing found"] else [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^ (if len <= lim then "" else " (" ^ string_of_int lim ^ " displayed)") ^ ":"), Pretty.str ""] @ map prt_fact thms) |> Pretty.chunks |> Pretty.writeln end;end;