| author | wenzelm | 
| Thu, 16 Mar 2023 15:16:17 +0100 | |
| changeset 77675 | 9e5f8f6e58a0 | 
| parent 74555 | 3ba399ecdfaf | 
| child 77854 | 64533f3818a4 | 
| permissions | -rw-r--r-- | 
| 30143 | 1 | (* Title: Pure/Tools/find_theorems.ML | 
| 26283 | 2 | Author: Rafal Kolanski and Gerwin Klein, NICTA | 
| 46718 | 3 | Author: Lars Noschinski and Alexander Krauss, TU Muenchen | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 4 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 5 | Retrieve theorems from proof context. | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 6 | *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 7 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 8 | signature FIND_THEOREMS = | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 9 | sig | 
| 16036 | 10 | datatype 'term criterion = | 
| 46717 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 wenzelm parents: 
46716diff
changeset | 11 | Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term | 
| 43070 | 12 |   type 'term query = {
 | 
| 13 | goal: thm option, | |
| 14 | limit: int option, | |
| 15 | rem_dups: bool, | |
| 16 | criteria: (bool * 'term criterion) list | |
| 17 | } | |
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 18 | val query_parser: (bool * string criterion) list parser | 
| 52925 | 19 | val read_query: Position.T -> string -> (bool * string criterion) list | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 20 | val find_theorems: Proof.context -> thm option -> int option -> bool -> | 
| 43067 | 21 | (bool * term criterion) list -> int option * (Facts.ref * thm) list | 
| 22 | val find_theorems_cmd: Proof.context -> thm option -> int option -> bool -> | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 23 | (bool * string criterion) list -> int option * (Facts.ref * thm) list | 
| 30186 
1f836e949ac2
replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
 wenzelm parents: 
30143diff
changeset | 24 | val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T | 
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 25 | val pretty_theorems: Proof.state -> | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 26 | int option -> bool -> (bool * string criterion) list -> Pretty.T | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 27 | val proof_state: Toplevel.state -> Proof.state | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 28 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 29 | |
| 33301 | 30 | structure Find_Theorems: FIND_THEOREMS = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 31 | struct | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 32 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 33 | (** search criteria **) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 34 | |
| 16036 | 35 | datatype 'term criterion = | 
| 46717 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 wenzelm parents: 
46716diff
changeset | 36 | Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term; | 
| 16036 | 37 | |
| 38 | fun read_criterion _ (Name name) = Name name | |
| 39 | | read_criterion _ Intro = Intro | |
| 40 | | read_criterion _ Elim = Elim | |
| 41 | | read_criterion _ Dest = Dest | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 42 | | read_criterion _ Solves = Solves | 
| 42360 | 43 | | read_criterion ctxt (Simp str) = Simp (Proof_Context.read_term_pattern ctxt str) | 
| 59097 
4b05ce4c84b0
revert "better" handling of abbreviation from c61fe520602b
 haftmann parents: 
59096diff
changeset | 44 | | read_criterion ctxt (Pattern str) = Pattern (Proof_Context.read_term_pattern ctxt str); | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 45 | |
| 16036 | 46 | fun pretty_criterion ctxt (b, c) = | 
| 47 | let | |
| 48 | fun prfx s = if b then s else "-" ^ s; | |
| 49 | in | |
| 50 | (case c of | |
| 51 | Name name => Pretty.str (prfx "name: " ^ quote name) | |
| 52 | | Intro => Pretty.str (prfx "intro") | |
| 53 | | Elim => Pretty.str (prfx "elim") | |
| 54 | | Dest => Pretty.str (prfx "dest") | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 55 | | Solves => Pretty.str (prfx "solves") | 
| 16088 | 56 | | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, | 
| 24920 | 57 | Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] | 
| 56914 | 58 | | Pattern pat => Pretty.enclose (prfx "\"") "\"" | 
| 24920 | 59 | [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) | 
| 16036 | 60 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 61 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 62 | |
| 43620 | 63 | |
| 43070 | 64 | (** queries **) | 
| 65 | ||
| 66 | type 'term query = {
 | |
| 67 | goal: thm option, | |
| 68 | limit: int option, | |
| 69 | rem_dups: bool, | |
| 70 | criteria: (bool * 'term criterion) list | |
| 71 | }; | |
| 72 | ||
| 73 | fun map_criteria f {goal, limit, rem_dups, criteria} =
 | |
| 46718 | 74 |   {goal = goal, limit = limit, rem_dups = rem_dups, criteria = f criteria};
 | 
| 43070 | 75 | |
| 43620 | 76 | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 77 | (** search criterion filters **) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 78 | |
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 79 | (*generated filters are to be of the form | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 80 | input: (Facts.ref * thm) | 
| 53632 | 81 | output: (p:int, s:int, t:int) option, where | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 82 | NONE indicates no match | 
| 17106 | 83 | p is the primary sorting criterion | 
| 53632 | 84 | (eg. size of term) | 
| 85 | s is the secondary sorting criterion | |
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 86 | (eg. number of assumptions in the theorem) | 
| 53632 | 87 | t is the tertiary sorting criterion | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 88 | (eg. size of the substitution for intro, elim and dest) | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 89 | when applying a set of filters to a thm, fold results in: | 
| 53632 | 90 | (max p, max s, sum of all t) | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 91 | *) | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 92 | |
| 16088 | 93 | |
| 94 | (* matching theorems *) | |
| 17106 | 95 | |
| 59970 | 96 | fun is_nontrivial ctxt = Term.is_Const o Term.head_of o Object_Logic.drop_judgment ctxt; | 
| 16088 | 97 | |
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 98 | (*extract terms from term_src, refine them to the parts that concern us, | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 99 | if po try match them against obj else vice versa. | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 100 | trivial matches are ignored. | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 101 | returns: smallest substitution size*) | 
| 46717 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 wenzelm parents: 
46716diff
changeset | 102 | fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src = | 
| 16088 | 103 | let | 
| 42360 | 104 | val thy = Proof_Context.theory_of ctxt; | 
| 16088 | 105 | |
| 16486 | 106 | fun matches pat = | 
| 59970 | 107 | is_nontrivial ctxt pat andalso | 
| 46717 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 wenzelm parents: 
46716diff
changeset | 108 | Pattern.matches thy (if po then (pat, obj) else (obj, pat)); | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 109 | |
| 52940 | 110 | fun subst_size pat = | 
| 18184 | 111 | let val (_, subst) = | 
| 112 | Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty) | |
| 17205 | 113 | in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end; | 
| 16088 | 114 | |
| 52941 | 115 | fun best_match [] = NONE | 
| 116 | | best_match xs = SOME (foldl1 Int.min xs); | |
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 117 | |
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 118 | val match_thm = matches o refine_term; | 
| 16486 | 119 | in | 
| 52940 | 120 | map (subst_size o refine_term) (filter match_thm (extract_terms term_src)) | 
| 52941 | 121 | |> best_match | 
| 16088 | 122 | end; | 
| 123 | ||
| 124 | ||
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 125 | (* filter_name *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 126 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 127 | fun filter_name str_pat (thmref, _) = | 
| 70545 | 128 | if match_string str_pat (Facts.ref_name thmref) | 
| 53632 | 129 | then SOME (0, 0, 0) else NONE; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 130 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 131 | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 132 | (* filter intro/elim/dest/solves rules *) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 133 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 134 | fun filter_dest ctxt goal (_, thm) = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 135 | let | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 136 | val extract_dest = | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 137 | (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm], | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 138 | hd o Logic.strip_imp_prems); | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 139 | val prems = Logic.prems_of_goal goal 1; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 140 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 141 | fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm; | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19476diff
changeset | 142 | val successful = prems |> map_filter try_subst; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 143 | in | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 144 | (*if possible, keep best substitution (one with smallest size)*) | 
| 17106 | 145 | (*dest rules always have assumptions, so a dest with one | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 146 | assumption is as good as an intro rule with none*) | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 147 | if not (null successful) then | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 148 | SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 149 | else NONE | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 150 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 151 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 152 | fun filter_intro ctxt goal (_, thm) = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 153 | let | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 154 | val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl); | 
| 16036 | 155 | val concl = Logic.concl_of_goal goal 1; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 156 | in | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 157 | (case is_matching_thm extract_intro ctxt true concl thm of | 
| 59096 | 158 | SOME k => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, k) | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 159 | | NONE => NONE) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 160 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 161 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 162 | fun filter_elim ctxt goal (_, thm) = | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 163 | if Thm.nprems_of thm > 0 then | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 164 | let | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 165 | val rule = Thm.full_prop_of thm; | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 166 | val prems = Logic.prems_of_goal goal 1; | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 167 | val goal_concl = Logic.concl_of_goal goal 1; | 
| 26283 | 168 | val rule_mp = hd (Logic.strip_imp_prems rule); | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 169 | val rule_concl = Logic.strip_imp_concl rule; | 
| 57690 | 170 |       fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);  (* FIXME ?!? *)
 | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 171 | val rule_tree = combine rule_mp rule_concl; | 
| 26283 | 172 | fun goal_tree prem = combine prem goal_concl; | 
| 46717 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 wenzelm parents: 
46716diff
changeset | 173 | fun try_subst prem = is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree; | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19476diff
changeset | 174 | val successful = prems |> map_filter try_subst; | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 175 | in | 
| 32798 | 176 | (*elim rules always have assumptions, so an elim with one | 
| 177 | assumption is as good as an intro rule with none*) | |
| 59970 | 178 | if is_nontrivial ctxt (Thm.major_prem_of thm) andalso not (null successful) then | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 179 | SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 180 | else NONE | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 181 | end | 
| 46718 | 182 | else NONE; | 
| 16036 | 183 | |
| 30143 | 184 | fun filter_solves ctxt goal = | 
| 185 | let | |
| 64983 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64556diff
changeset | 186 | val thy' = Proof_Context.theory_of ctxt | 
| 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64556diff
changeset | 187 | |> Context_Position.set_visible_global false; | 
| 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64556diff
changeset | 188 | val ctxt' = Proof_Context.transfer thy' ctxt | 
| 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64556diff
changeset | 189 | |> Context_Position.set_visible false; | 
| 52704 
b824497c8e86
modify background theory where it is actually required (cf. 51dfdcd88e84);
 wenzelm parents: 
52703diff
changeset | 190 | val goal' = Thm.transfer thy' goal; | 
| 
b824497c8e86
modify background theory where it is actually required (cf. 51dfdcd88e84);
 wenzelm parents: 
52703diff
changeset | 191 | |
| 52941 | 192 | fun limited_etac thm i = | 
| 67147 | 193 | Seq.take (Options.default_int \<^system_option>\<open>find_theorems_tactic_limit\<close>) o | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59098diff
changeset | 194 | eresolve_tac ctxt' [thm] i; | 
| 30143 | 195 | fun try_thm thm = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59098diff
changeset | 196 | if Thm.no_prems thm then resolve_tac ctxt' [thm] 1 goal' | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53633diff
changeset | 197 | else | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53633diff
changeset | 198 | (limited_etac thm THEN_ALL_NEW (Goal.norm_hhf_tac ctxt' THEN' Method.assm_tac ctxt')) | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53633diff
changeset | 199 | 1 goal'; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 200 | in | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 201 | fn (_, thm) => | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 202 | if is_some (Seq.pull (try_thm thm)) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 203 | then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 204 | else NONE | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 205 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 206 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 207 | |
| 16074 | 208 | (* filter_simp *) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 209 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 210 | fun filter_simp ctxt t (_, thm) = | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 211 | let | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 212 | val mksimps = Simplifier.mksimps ctxt; | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 213 | val extract_simp = | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 214 | (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl); | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 215 | in | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 216 | (case is_matching_thm extract_simp ctxt false t thm of | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 217 | SOME ss => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, ss) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 218 | | NONE => NONE) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 219 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 220 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 221 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 222 | (* filter_pattern *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 223 | |
| 59098 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 224 | fun expand_abs t = | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 225 | let | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 226 | val m = Term.maxidx_of_term t + 1; | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 227 | val vs = strip_abs_vars t; | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 228 | val ts = map_index (fn (k, (_, T)) => Var ((Name.aT, m + k), T)) vs; | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 229 | in betapplys (t, ts) end; | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 230 | |
| 32798 | 231 | fun get_names t = Term.add_const_names t (Term.add_free_names t []); | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 232 | |
| 59095 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 haftmann parents: 
59083diff
changeset | 233 | (* Does pat match a subterm of obj? *) | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 234 | fun matches_subterm ctxt (pat, obj) = | 
| 59095 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 haftmann parents: 
59083diff
changeset | 235 | let | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 236 | val thy = Proof_Context.theory_of ctxt; | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 237 | fun matches obj ctxt' = Pattern.matches thy (pat, obj) orelse | 
| 59095 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 haftmann parents: 
59083diff
changeset | 238 | (case obj of | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 239 | Abs (_, T, t) => | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 240 | let val ((_, t'), ctxt'') = Variable.dest_abs obj ctxt' | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 241 | in matches t' ctxt'' end | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 242 | | t $ u => matches t ctxt' orelse matches u ctxt' | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 243 | | _ => false); | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 244 | in matches obj ctxt end; | 
| 59095 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 haftmann parents: 
59083diff
changeset | 245 | |
| 52940 | 246 | (*Including all constants and frees is only sound because matching | 
| 247 | uses higher-order patterns. If full matching were used, then | |
| 248 | constants that may be subject to beta-reduction after substitution | |
| 249 | of frees should not be included for LHS set because they could be | |
| 250 | thrown away by the substituted function. E.g. for (?F 1 2) do not | |
| 67721 | 251 | include 1 or 2, if it were possible for ?F to be (\<lambda>x y. 3). The | 
| 52940 | 252 | largest possible set should always be included on the RHS.*) | 
| 30143 | 253 | |
| 254 | fun filter_pattern ctxt pat = | |
| 255 | let | |
| 59098 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 256 | val pat' = (expand_abs o Envir.eta_contract) pat; | 
| 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 haftmann parents: 
59097diff
changeset | 257 | val pat_consts = get_names pat'; | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 258 | fun check ((x, thm), NONE) = check ((x, thm), SOME (get_names (Thm.full_prop_of thm))) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 259 | | check ((_, thm), c as SOME thm_consts) = | 
| 33038 | 260 | (if subset (op =) (pat_consts, thm_consts) andalso | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
70545diff
changeset | 261 | matches_subterm ctxt (pat', Thm.full_prop_of thm) | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 262 | then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0) else NONE, c); | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 263 | in check end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 264 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 265 | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 266 | (* interpret criteria as filters *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 267 | |
| 16036 | 268 | local | 
| 269 | ||
| 270 | fun err_no_goal c = | |
| 271 |   error ("Current goal required for " ^ c ^ " search criterion");
 | |
| 272 | ||
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 273 | fun filter_crit _ _ (Name name) = apfst (filter_name name) | 
| 16036 | 274 | | filter_crit _ NONE Intro = err_no_goal "intro" | 
| 275 | | filter_crit _ NONE Elim = err_no_goal "elim" | |
| 276 | | filter_crit _ NONE Dest = err_no_goal "dest" | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 277 | | filter_crit _ NONE Solves = err_no_goal "solves" | 
| 52940 | 278 | | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt (Thm.prop_of goal)) | 
| 279 | | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (Thm.prop_of goal)) | |
| 280 | | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (Thm.prop_of goal)) | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 281 | | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal) | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 282 | | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat) | 
| 16088 | 283 | | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat; | 
| 16036 | 284 | |
| 53632 | 285 | fun opt_not x = if is_some x then NONE else SOME (0, 0, 0); | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 286 | |
| 53632 | 287 | fun opt_add (SOME (a, c, x)) (SOME (b, d, y)) = SOME (Int.max (a,b), Int.max (c, d), x + y : int) | 
| 26283 | 288 | | opt_add _ _ = NONE; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 289 | |
| 30143 | 290 | fun app_filters thm = | 
| 291 | let | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 292 | fun app (NONE, _, _) = NONE | 
| 32798 | 293 | | app (SOME v, _, []) = SOME (v, thm) | 
| 30143 | 294 | | app (r, consts, f :: fs) = | 
| 295 | let val (r', consts') = f (thm, consts) | |
| 296 | in app (opt_add r r', consts', fs) end; | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 297 | in app end; | 
| 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 298 | |
| 16036 | 299 | in | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 300 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 301 | fun filter_criterion ctxt opt_goal (b, c) = | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 302 | (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 303 | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 304 | fun sorted_filter filters thms = | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 305 | let | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 306 | fun eval_filters thm = app_filters thm (SOME (0, 0, 0), NONE, filters); | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 307 | |
| 53632 | 308 | (*filters return: (thm size, number of assumptions, substitution size) option, so | 
| 309 | sort according to size of thm first, then number of assumptions, | |
| 310 | then by the substitution size, then by term order *) | |
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 311 | fun result_ord (((p0, s0, t0), (_, thm0)), ((p1, s1, t1), (_, thm1))) = | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 312 | prod_ord int_ord (prod_ord int_ord (prod_ord int_ord Term_Ord.term_ord)) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 313 | ((p1, (s1, (t1, Thm.full_prop_of thm1))), (p0, (s0, (t0, Thm.full_prop_of thm0)))); | 
| 46977 | 314 | in | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 315 | grouped 100 Par_List.map eval_filters thms | 
| 46977 | 316 | |> map_filter I |> sort result_ord |> map #2 | 
| 317 | end; | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 318 | |
| 30822 | 319 | fun lazy_filter filters = | 
| 320 | let | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 321 | fun lazy_match thms = Seq.make (fn () => first_match thms) | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 322 | and first_match [] = NONE | 
| 30822 | 323 | | first_match (thm :: thms) = | 
| 53632 | 324 | (case app_filters thm (SOME (0, 0, 0), NONE, filters) of | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 325 | NONE => first_match thms | 
| 30822 | 326 | | SOME (_, t) => SOME (t, lazy_match thms)); | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 327 | in lazy_match end; | 
| 30822 | 328 | |
| 16036 | 329 | end; | 
| 330 | ||
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 331 | |
| 52940 | 332 | (* removing duplicates, preferring nicer names, roughly O(n log n) *) | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 333 | |
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 334 | local | 
| 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 335 | |
| 27486 | 336 | val index_ord = option_ord (K EQUAL); | 
| 59916 | 337 | val hidden_ord = bool_ord o apply2 Long_Name.is_hidden; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 338 | val qual_ord = int_ord o apply2 Long_Name.qualification; | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 339 | val txt_ord = int_ord o apply2 size; | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 340 | |
| 63080 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 341 | fun nicer_name ((a, x), i) ((b, y), j) = | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 342 | (case bool_ord (a, b) of EQUAL => | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 343 | (case hidden_ord (x, y) of EQUAL => | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 344 | (case index_ord (i, j) of EQUAL => | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 345 | (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord) | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 346 | | ord => ord) | 
| 27486 | 347 | | ord => ord) | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 348 | | ord => ord) <> GREATER; | 
| 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 349 | |
| 29848 | 350 | fun rem_cdups nicer xs = | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 351 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 352 | fun rem_c rev_seen [] = rev rev_seen | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 353 | | rem_c rev_seen [x] = rem_c (x :: rev_seen) [] | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 354 | | rem_c rev_seen ((x as ((n, thm), _)) :: (y as ((n', thm'), _)) :: rest) = | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 355 | if Thm.eq_thm_prop (thm, thm') | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 356 | then rem_c rev_seen ((if nicer n n' then x else y) :: rest) | 
| 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 357 | else rem_c (x :: rev_seen) (y :: rest); | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 358 | in rem_c [] xs end; | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 359 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 360 | in | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 361 | |
| 30143 | 362 | fun nicer_shortest ctxt = | 
| 363 | let | |
| 56143 
ed2b660a52a1
more accurate resolution of hybrid facts, which actually changes the sort order of results;
 wenzelm parents: 
56141diff
changeset | 364 | fun extern_shortest name = | 
| 63080 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 365 | let | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 366 | val facts = Proof_Context.facts_of_fact ctxt name; | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 367 | val space = Facts.space_of facts; | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 368 | in (Facts.is_dynamic facts name, Name_Space.extern_shortest ctxt space name) end; | 
| 29848 | 369 | |
| 370 | fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) = | |
| 55672 
5e25cc741ab9
support for completion within the formal context;
 wenzelm parents: 
55671diff
changeset | 371 | nicer_name (extern_shortest x, i) (extern_shortest y, j) | 
| 29848 | 372 | | nicer (Facts.Fact _) (Facts.Named _) = true | 
| 55670 | 373 | | nicer (Facts.Named _) (Facts.Fact _) = false | 
| 374 | | nicer (Facts.Fact _) (Facts.Fact _) = true; | |
| 29848 | 375 | in nicer end; | 
| 376 | ||
| 377 | fun rem_thm_dups nicer xs = | |
| 52940 | 378 | (xs ~~ (1 upto length xs)) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 379 | |> sort (Term_Ord.fast_term_ord o apply2 (Thm.full_prop_of o #2 o #1)) | 
| 29848 | 380 | |> rem_cdups nicer | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 381 | |> sort (int_ord o apply2 #2) | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 382 | |> map #1; | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 383 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 384 | end; | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 385 | |
| 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 386 | |
| 52941 | 387 | |
| 388 | (** main operations **) | |
| 389 | ||
| 390 | (* filter_theorems *) | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 391 | |
| 26283 | 392 | fun all_facts_of ctxt = | 
| 33381 | 393 | let | 
| 61054 | 394 | val thy = Proof_Context.theory_of ctxt; | 
| 395 | val transfer = Global_Theory.transfer_theories thy; | |
| 56158 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 wenzelm parents: 
56143diff
changeset | 396 | val local_facts = Proof_Context.facts_of ctxt; | 
| 61054 | 397 | val global_facts = Global_Theory.facts_of thy; | 
| 56141 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 wenzelm parents: 
56140diff
changeset | 398 | in | 
| 63080 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 399 | (Facts.dest_all (Context.Proof ctxt) false [global_facts] local_facts @ | 
| 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 400 | Facts.dest_all (Context.Proof ctxt) false [] global_facts) | 
| 61054 | 401 | |> maps Facts.selections | 
| 402 | |> map (apsnd transfer) | |
| 56141 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 wenzelm parents: 
56140diff
changeset | 403 | end; | 
| 17972 | 404 | |
| 43070 | 405 | fun filter_theorems ctxt theorems query = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 406 | let | 
| 46718 | 407 |     val {goal = opt_goal, limit = opt_limit, rem_dups, criteria} = query;
 | 
| 43069 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 408 | val filters = map (filter_criterion ctxt opt_goal) criteria; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 409 | |
| 41844 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 krauss parents: 
41841diff
changeset | 410 | fun find_all theorems = | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 411 | let | 
| 41844 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 krauss parents: 
41841diff
changeset | 412 | val raw_matches = sorted_filter filters theorems; | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 413 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 414 | val matches = | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 415 | if rem_dups | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 416 | then rem_thm_dups (nicer_shortest ctxt) raw_matches | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 417 | else raw_matches; | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 418 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 419 | val len = length matches; | 
| 67147 | 420 | val lim = the_default (Options.default_int \<^system_option>\<open>find_theorems_limit\<close>) opt_limit; | 
| 34088 | 421 | in (SOME len, drop (Int.max (len - lim, 0)) matches) end; | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 422 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 423 | val find = | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 424 | if rem_dups orelse is_none opt_limit | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 425 | then find_all | 
| 30822 | 426 | else pair NONE o Seq.list_of o Seq.take (the opt_limit) o lazy_filter filters; | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 427 | |
| 41844 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 krauss parents: 
41841diff
changeset | 428 | in find theorems end; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 429 | |
| 46718 | 430 | fun filter_theorems_cmd ctxt theorems raw_query = | 
| 52941 | 431 | filter_theorems ctxt theorems (map_criteria (map (apsnd (read_criterion ctxt))) raw_query); | 
| 432 | ||
| 433 | ||
| 434 | (* find_theorems *) | |
| 435 | ||
| 436 | local | |
| 43067 | 437 | |
| 438 | fun gen_find_theorems filter ctxt opt_goal opt_limit rem_dups raw_criteria = | |
| 43069 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 439 | let | 
| 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 440 | val assms = | 
| 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 441 | Proof_Context.get_fact ctxt (Facts.named "local.assms") | 
| 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 442 | handle ERROR _ => []; | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61268diff
changeset | 443 | val add_prems = Seq.hd o TRY (Method.insert_tac ctxt assms 1); | 
| 43069 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 444 | val opt_goal' = Option.map add_prems opt_goal; | 
| 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 445 | in | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 446 | filter ctxt (all_facts_of ctxt) | 
| 46718 | 447 |       {goal = opt_goal', limit = opt_limit, rem_dups = rem_dups, criteria = raw_criteria}
 | 
| 43069 
88e45168272c
moved questionable goal modification out of filter_theorems
 krauss parents: 
43068diff
changeset | 448 | end; | 
| 30186 
1f836e949ac2
replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
 wenzelm parents: 
30143diff
changeset | 449 | |
| 52941 | 450 | in | 
| 451 | ||
| 43067 | 452 | val find_theorems = gen_find_theorems filter_theorems; | 
| 453 | val find_theorems_cmd = gen_find_theorems filter_theorems_cmd; | |
| 454 | ||
| 52941 | 455 | end; | 
| 456 | ||
| 457 | ||
| 458 | (* pretty_theorems *) | |
| 459 | ||
| 460 | local | |
| 461 | ||
| 49888 | 462 | fun pretty_ref ctxt thmref = | 
| 463 | let | |
| 464 | val (name, sel) = | |
| 465 | (case thmref of | |
| 466 | Facts.Named ((name, _), sel) => (name, sel) | |
| 467 | | Facts.Fact _ => raise Fail "Illegal literal fact"); | |
| 468 | in | |
| 63080 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 wenzelm parents: 
62969diff
changeset | 469 | [Pretty.marks_str (#1 (Proof_Context.markup_extern_fact ctxt name), name), | 
| 56141 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 wenzelm parents: 
56140diff
changeset | 470 | Pretty.str (Facts.string_of_selection sel), Pretty.str ":", Pretty.brk 1] | 
| 49888 | 471 | end; | 
| 472 | ||
| 52941 | 473 | in | 
| 474 | ||
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 475 | fun pretty_thm ctxt (thmref, thm) = | 
| 61268 | 476 | Pretty.block (pretty_ref ctxt thmref @ [Thm.pretty_thm ctxt thm]); | 
| 41845 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 krauss parents: 
41844diff
changeset | 477 | |
| 52941 | 478 | fun pretty_theorems state opt_limit rem_dups raw_criteria = | 
| 30143 | 479 | let | 
| 52941 | 480 | val ctxt = Proof.context_of state; | 
| 64984 | 481 | val opt_goal = try (#goal o Proof.simple_goal) state; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 482 | val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; | 
| 52941 | 483 | |
| 52940 | 484 | val (opt_found, theorems) = | 
| 55671 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 wenzelm parents: 
55670diff
changeset | 485 | filter_theorems ctxt (all_facts_of ctxt) | 
| 52855 | 486 |         {goal = opt_goal, limit = opt_limit, rem_dups = rem_dups, criteria = criteria};
 | 
| 41845 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 krauss parents: 
41844diff
changeset | 487 | val returned = length theorems; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
31042diff
changeset | 488 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 489 | val tally_msg = | 
| 52940 | 490 | (case opt_found of | 
| 38335 | 491 | NONE => "displaying " ^ string_of_int returned ^ " theorem(s)" | 
| 30822 | 492 | | SOME found => | 
| 38335 | 493 | "found " ^ string_of_int found ^ " theorem(s)" ^ | 
| 30822 | 494 | (if returned < found | 
| 495 |              then " (" ^ string_of_int returned ^ " displayed)"
 | |
| 496 | else "")); | |
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 497 | val position_markup = Position.markup (Position.thread_data ()) Markup.position; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 498 | in | 
| 56891 
48899c43b07d
tuned message -- more context for detached window etc.;
 wenzelm parents: 
56621diff
changeset | 499 | Pretty.block | 
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 500 | (Pretty.fbreaks | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 501 | (Pretty.mark position_markup (Pretty.keyword1 "find_theorems") :: | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 502 | map (pretty_criterion ctxt) criteria)) :: | 
| 38335 | 503 | Pretty.str "" :: | 
| 56908 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 504 | (if null theorems then [Pretty.str "found nothing"] | 
| 38335 | 505 | else | 
| 56908 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 506 | Pretty.str (tally_msg ^ ":") :: | 
| 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 507 | grouped 10 Par_List.map (Pretty.item o single o pretty_thm ctxt) (rev theorems)) | 
| 52855 | 508 | end |> Pretty.fbreaks |> curry Pretty.blk 0; | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 509 | |
| 52941 | 510 | end; | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 511 | |
| 32798 | 512 | |
| 46718 | 513 | |
| 52865 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 wenzelm parents: 
52863diff
changeset | 514 | (** Isar command syntax **) | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 515 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 516 | local | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 517 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 518 | val criterion = | 
| 62969 | 519 | Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.name) >> Name || | 
| 36950 | 520 | Parse.reserved "intro" >> K Intro || | 
| 521 | Parse.reserved "elim" >> K Elim || | |
| 522 | Parse.reserved "dest" >> K Dest || | |
| 523 | Parse.reserved "solves" >> K Solves || | |
| 524 | Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp || | |
| 525 | Parse.term >> Pattern; | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 526 | |
| 74555 | 527 | val query_keywords = Keyword.add_minor_keywords [":"] Keyword.empty_keywords; | 
| 58905 | 528 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 529 | in | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 530 | |
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 531 | val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 532 | |
| 52925 | 533 | fun read_query pos str = | 
| 59083 | 534 | Token.explode query_keywords pos str | 
| 52855 | 535 | |> filter Token.is_proper | 
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 536 | |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof))) | 
| 52925 | 537 | |> #1; | 
| 43068 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 krauss parents: 
43067diff
changeset | 538 | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 539 | end; | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 540 | |
| 52851 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 wenzelm parents: 
52788diff
changeset | 541 | |
| 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 wenzelm parents: 
52788diff
changeset | 542 | |
| 52865 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 wenzelm parents: 
52863diff
changeset | 543 | (** PIDE query operation **) | 
| 52854 
92932931bd82
more general Output.result: allow to update arbitrary properties;
 wenzelm parents: 
52851diff
changeset | 544 | |
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 545 | fun proof_state st = | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 546 | (case try Toplevel.proof_of st of | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 547 | SOME state => state | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 548 | | NONE => Proof.init (Toplevel.context_of st)); | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61841diff
changeset | 549 | |
| 52865 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 wenzelm parents: 
52863diff
changeset | 550 | val _ = | 
| 60610 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 551 |   Query_Operation.register {name = "find_theorems", pri = Task_Queue.urgent_pri}
 | 
| 61223 | 552 |     (fn {state = st, args, writeln_result, ...} =>
 | 
| 60610 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 553 | if can Toplevel.context_of st then | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 554 | let | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 555 | val [limit_arg, allow_dups_arg, query_arg] = args; | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 556 | val state = proof_state st; | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 557 | val opt_limit = Int.fromString limit_arg; | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 558 | val rem_dups = allow_dups_arg = "false"; | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 559 | val criteria = read_query Position.none query_arg; | 
| 61223 | 560 | in writeln_result (Pretty.string_of (pretty_theorems state opt_limit rem_dups criteria)) end | 
| 60610 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59970diff
changeset | 561 | else error "Unknown context"); | 
| 52851 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 wenzelm parents: 
52788diff
changeset | 562 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 563 | end; |