| author | wenzelm | 
| Thu, 20 May 2010 16:25:22 +0200 | |
| changeset 37011 | f692d6178e4e | 
| parent 36953 | 2af1ad9aa1a3 | 
| child 38334 | c677c2c1d333 | 
| permissions | -rw-r--r-- | 
| 30143 | 1 | (* Title: Pure/Tools/find_theorems.ML | 
| 26283 | 2 | Author: Rafal Kolanski and Gerwin Klein, NICTA | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 3 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 4 | Retrieve theorems from proof context. | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 5 | *) | 
| 
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 | signature FIND_THEOREMS = | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 8 | sig | 
| 16036 | 9 | datatype 'term criterion = | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 10 | Name of string | Intro | IntroIff | Elim | Dest | Solves | Simp of 'term | | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 11 | Pattern of 'term | 
| 32738 | 12 | val tac_limit: int Unsynchronized.ref | 
| 13 | val limit: int Unsynchronized.ref | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 14 | val find_theorems: Proof.context -> thm option -> int option -> bool -> | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 15 | (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 | 16 | val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 17 | val print_theorems: Proof.context -> thm option -> int option -> bool -> | 
| 16036 | 18 | (bool * string criterion) list -> unit | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 19 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 20 | |
| 33301 | 21 | structure Find_Theorems: FIND_THEOREMS = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 22 | struct | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 23 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 24 | (** search criteria **) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 25 | |
| 16036 | 26 | datatype 'term criterion = | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 27 | Name of string | Intro | IntroIff | Elim | Dest | Solves | Simp of 'term | | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 28 | Pattern of 'term; | 
| 16036 | 29 | |
| 33036 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 30 | fun apply_dummies tm = | 
| 33301 | 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; | |
| 33036 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 35 | |
| 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 36 | fun parse_pattern ctxt nm = | 
| 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 37 | let | 
| 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 38 | val consts = ProofContext.consts_of ctxt; | 
| 33301 | 39 | val nm' = | 
| 40 | (case Syntax.parse_term ctxt nm of | |
| 41 | Const (c, _) => c | |
| 42 | | _ => Consts.intern consts nm); | |
| 33036 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 43 | in | 
| 33301 | 44 | (case try (Consts.the_abbreviation consts) nm' of | 
| 33923 
7fc1ab75b4fa
Expand nested abbreviations before applying dummy patterns.
 kleing parents: 
33382diff
changeset | 45 | SOME (_, rhs) => apply_dummies (ProofContext.expand_abbrevs ctxt rhs) | 
| 33301 | 46 | | NONE => ProofContext.read_term_pattern ctxt nm) | 
| 33036 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 47 | end; | 
| 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 48 | |
| 16036 | 49 | fun read_criterion _ (Name name) = Name name | 
| 50 | | read_criterion _ Intro = Intro | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 51 | | read_criterion _ IntroIff = IntroIff | 
| 16036 | 52 | | read_criterion _ Elim = Elim | 
| 53 | | read_criterion _ Dest = Dest | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 54 | | read_criterion _ Solves = Solves | 
| 24683 | 55 | | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str) | 
| 33036 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 kleing parents: 
33029diff
changeset | 56 | | read_criterion ctxt (Pattern str) = Pattern (parse_pattern ctxt str); | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 57 | |
| 16036 | 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") | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 65 | | IntroIff => Pretty.str (prfx "introiff") | 
| 16036 | 66 | | Elim => Pretty.str (prfx "elim") | 
| 67 | | Dest => Pretty.str (prfx "dest") | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 68 | | Solves => Pretty.str (prfx "solves") | 
| 16088 | 69 | | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, | 
| 24920 | 70 | Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] | 
| 16036 | 71 | | Pattern pat => Pretty.enclose (prfx " \"") "\"" | 
| 24920 | 72 | [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) | 
| 16036 | 73 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 74 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 75 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 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 | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 80 | input: (Facts.ref * thm) | 
| 17106 | 81 | output: (p:int, s: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 | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 84 | (eg. number of assumptions in the theorem) | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 85 | s is the secondary sorting criterion | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 86 | (eg. size of the substitution for intro, elim and dest) | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 87 | when applying a set of filters to a thm, fold results in: | 
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 88 | (biggest p, sum of all s) | 
| 17106 | 89 | currently p and s only matter for intro, elim, dest and simp filters, | 
| 90 | otherwise the default ordering is used. | |
| 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 | |
| 35625 | 96 | fun is_nontrivial thy = Term.is_Const o Term.head_of o Object_Logic.drop_judgment thy; | 
| 16088 | 97 | |
| 32798 | 98 | (*educated guesses on HOL*) (* FIXME broken *) | 
| 99 | val boolT = Type ("bool", []);
 | |
| 100 | val iff_const = Const ("op =", boolT --> boolT --> boolT);
 | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 101 | |
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 102 | (*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 | 103 | 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 | 104 | trivial matches are ignored. | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 105 | returns: smallest substitution size*) | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 106 | fun is_matching_thm doiff (extract_terms, refine_term) ctxt po obj term_src = | 
| 16088 | 107 | let | 
| 17106 | 108 | val thy = ProofContext.theory_of ctxt; | 
| 16088 | 109 | |
| 32798 | 110 | fun check_match pat = Pattern.matches thy (if po then (pat, obj) else (obj, pat)); | 
| 16486 | 111 | fun matches pat = | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 112 | let | 
| 35625 | 113 | val jpat = Object_Logic.drop_judgment thy pat; | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 114 | val c = Term.head_of jpat; | 
| 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 115 | val pats = | 
| 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 116 | if Term.is_Const c | 
| 32798 | 117 | then | 
| 118 | if doiff andalso c = iff_const then | |
| 35625 | 119 | (pat :: map (Object_Logic.ensure_propT thy) (snd (strip_comb jpat))) | 
| 32798 | 120 | |> filter (is_nontrivial thy) | 
| 121 | else [pat] | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 122 | else []; | 
| 32798 | 123 | in filter check_match pats end; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 124 | |
| 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 125 | fun substsize pat = | 
| 18184 | 126 | let val (_, subst) = | 
| 127 | Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty) | |
| 17205 | 128 | in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end; | 
| 16088 | 129 | |
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 130 | fun bestmatch [] = NONE | 
| 33029 | 131 | | bestmatch xs = SOME (foldl1 Int.min xs); | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 132 | |
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 133 | val match_thm = matches o refine_term; | 
| 16486 | 134 | in | 
| 32798 | 135 | maps match_thm (extract_terms term_src) | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 136 | |> map substsize | 
| 26283 | 137 | |> bestmatch | 
| 16088 | 138 | end; | 
| 139 | ||
| 140 | ||
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 141 | (* filter_name *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 142 | |
| 17106 | 143 | fun filter_name str_pat (thmref, _) = | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 144 | if match_string str_pat (Facts.name_of_ref thmref) | 
| 17205 | 145 | then SOME (0, 0) else NONE; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 146 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 147 | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 148 | (* filter intro/elim/dest/solves rules *) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 149 | |
| 17205 | 150 | fun filter_dest ctxt goal (_, thm) = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 151 | let | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 152 | val extract_dest = | 
| 17205 | 153 | (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 | 154 | hd o Logic.strip_imp_prems); | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 155 | val prems = Logic.prems_of_goal goal 1; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 156 | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 157 | fun try_subst prem = is_matching_thm false extract_dest ctxt true prem thm; | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19476diff
changeset | 158 | val successful = prems |> map_filter try_subst; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 159 | in | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 160 | (*if possible, keep best substitution (one with smallest size)*) | 
| 17106 | 161 | (*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 | 162 | assumption is as good as an intro rule with none*) | 
| 17205 | 163 | if not (null successful) | 
| 33029 | 164 | then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 165 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 166 | |
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 167 | fun filter_intro doiff ctxt goal (_, thm) = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 168 | let | 
| 17205 | 169 | val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl); | 
| 16036 | 170 | val concl = Logic.concl_of_goal goal 1; | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 171 | val ss = is_matching_thm doiff extract_intro ctxt true concl thm; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 172 | in | 
| 18939 | 173 | if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 174 | end; | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 175 | |
| 17205 | 176 | fun filter_elim ctxt goal (_, thm) = | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 177 | if not (Thm.no_prems thm) then | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 178 | let | 
| 17205 | 179 | 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 | 180 | 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 | 181 | val goal_concl = Logic.concl_of_goal goal 1; | 
| 26283 | 182 | 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 | 183 | val rule_concl = Logic.strip_imp_concl rule; | 
| 26283 | 184 |       fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);
 | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 185 | val rule_tree = combine rule_mp rule_concl; | 
| 26283 | 186 | fun goal_tree prem = combine prem goal_concl; | 
| 17106 | 187 | fun try_subst prem = | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 188 | is_matching_thm false (single, I) ctxt true (goal_tree prem) rule_tree; | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19476diff
changeset | 189 | 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 | 190 | in | 
| 32798 | 191 | (*elim rules always have assumptions, so an elim with one | 
| 192 | assumption is as good as an intro rule with none*) | |
| 17106 | 193 | if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm) | 
| 17205 | 194 | andalso not (null successful) | 
| 33029 | 195 | then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 196 | end | 
| 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 197 | else NONE | 
| 16036 | 198 | |
| 32738 | 199 | val tac_limit = Unsynchronized.ref 5; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 200 | |
| 30143 | 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 | |
| 30318 
3d03190d2864
replaced archaic use of rep_ss by Simplifier.mksimps;
 wenzelm parents: 
30234diff
changeset | 206 | else (etacn thm THEN_ALL_NEW (Goal.norm_hhf_tac THEN' Method.assm_tac ctxt)) 1 goal; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 207 | in | 
| 30143 | 208 | fn (_, thm) => | 
| 32798 | 209 | if is_some (Seq.pull (try_thm thm)) | 
| 30143 | 210 | then SOME (Thm.nprems_of thm, 0) else NONE | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 211 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 212 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 213 | |
| 16074 | 214 | (* filter_simp *) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 215 | |
| 17205 | 216 | fun filter_simp ctxt t (_, thm) = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 217 | let | 
| 32149 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
32091diff
changeset | 218 | val mksimps = Simplifier.mksimps (simpset_of ctxt); | 
| 17106 | 219 | val extract_simp = | 
| 30318 
3d03190d2864
replaced archaic use of rep_ss by Simplifier.mksimps;
 wenzelm parents: 
30234diff
changeset | 220 | (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl); | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 221 | val ss = is_matching_thm false extract_simp ctxt false t thm; | 
| 17106 | 222 | in | 
| 18939 | 223 | if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 224 | end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 225 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 226 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 227 | (* filter_pattern *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 228 | |
| 32798 | 229 | fun get_names t = Term.add_const_names t (Term.add_free_names t []); | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 230 | fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm); | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 231 | |
| 30143 | 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 | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 245 | val pat_consts = get_names pat; | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 246 | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 247 | fun check (t, NONE) = check (t, SOME (get_thm_names t)) | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 248 | | check ((_, thm), c as SOME thm_consts) = | 
| 33038 | 249 | (if subset (op =) (pat_consts, thm_consts) andalso | 
| 32798 | 250 | Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm) | 
| 251 | then SOME (0, 0) else NONE, c); | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 252 | in check end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 253 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 254 | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 255 | (* interpret criteria as filters *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 256 | |
| 16036 | 257 | local | 
| 258 | ||
| 259 | fun err_no_goal c = | |
| 260 |   error ("Current goal required for " ^ c ^ " search criterion");
 | |
| 261 | ||
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 262 | val fix_goal = Thm.prop_of; | 
| 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 263 | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 264 | fun filter_crit _ _ (Name name) = apfst (filter_name name) | 
| 16036 | 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" | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 268 | | filter_crit _ NONE Solves = err_no_goal "solves" | 
| 31042 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 269 | | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro false ctxt (fix_goal goal)) | 
| 
d452117ba564
Prototype introiff option for find_theorems.
 Timothy Bourke parents: 
30822diff
changeset | 270 | | filter_crit ctxt (SOME goal) IntroIff = apfst (filter_intro true ctxt (fix_goal goal)) | 
| 30143 | 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)) | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 273 | | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal) | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 274 | | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat) | 
| 16088 | 275 | | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat; | 
| 16036 | 276 | |
| 19502 | 277 | fun opt_not x = if is_some x then NONE else SOME (0, 0); | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 278 | |
| 17756 | 279 | fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int) | 
| 26283 | 280 | | opt_add _ _ = NONE; | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 281 | |
| 30143 | 282 | fun app_filters thm = | 
| 283 | let | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 284 | fun app (NONE, _, _) = NONE | 
| 32798 | 285 | | app (SOME v, _, []) = SOME (v, thm) | 
| 30143 | 286 | | app (r, consts, f :: fs) = | 
| 287 | let val (r', consts') = f (thm, consts) | |
| 288 | in app (opt_add r r', consts', fs) end; | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 289 | in app end; | 
| 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 290 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
31042diff
changeset | 291 | |
| 16036 | 292 | in | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 293 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 294 | fun filter_criterion ctxt opt_goal (b, c) = | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 295 | (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 | 296 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 297 | fun sorted_filter filters thms = | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 298 | let | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 299 | fun eval_filters thm = app_filters thm (SOME (0, 0), NONE, filters); | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 300 | |
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 301 | (*filters return: (number of assumptions, substitution size) option, so | 
| 16964 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 kleing parents: 
16895diff
changeset | 302 | sort (desc. in both cases) according to number of assumptions first, | 
| 16895 
df67fc190e06
Sort search results in order of relevance, where relevance =
 kleing parents: 
16486diff
changeset | 303 | then by the substitution size*) | 
| 17205 | 304 | fun thm_ord (((p0, s0), _), ((p1, s1), _)) = | 
| 305 | prod_ord int_ord int_ord ((p1, s1), (p0, s0)); | |
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 306 | in map_filter eval_filters thms |> sort thm_ord |> map #2 end; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 307 | |
| 30822 | 308 | fun lazy_filter filters = | 
| 309 | let | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 310 | 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 | 311 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 312 | and first_match [] = NONE | 
| 30822 | 313 | | first_match (thm :: thms) = | 
| 314 | (case app_filters thm (SOME (0, 0), NONE, filters) of | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 315 | NONE => first_match thms | 
| 30822 | 316 | | SOME (_, t) => SOME (t, lazy_match thms)); | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 317 | in lazy_match end; | 
| 30822 | 318 | |
| 16036 | 319 | end; | 
| 320 | ||
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 321 | |
| 22414 | 322 | (* removing duplicates, preferring nicer names, roughly n log n *) | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 323 | |
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 324 | local | 
| 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 325 | |
| 27486 | 326 | val index_ord = option_ord (K EQUAL); | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33039diff
changeset | 327 | val hidden_ord = bool_ord o pairself Name_Space.is_hidden; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30318diff
changeset | 328 | val qual_ord = int_ord o pairself (length o Long_Name.explode); | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 329 | val txt_ord = int_ord o pairself size; | 
| 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 330 | |
| 27486 | 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) | |
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 336 | | ord => ord) <> GREATER; | 
| 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 337 | |
| 29848 | 338 | fun rem_cdups nicer xs = | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 339 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 340 | fun rem_c rev_seen [] = rev rev_seen | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 341 | | rem_c rev_seen [x] = rem_c (x :: rev_seen) [] | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 342 | | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) = | 
| 30822 | 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) | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 346 | in rem_c [] xs end; | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 347 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 348 | in | 
| 25226 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 wenzelm parents: 
24920diff
changeset | 349 | |
| 30143 | 350 | fun nicer_shortest ctxt = | 
| 351 | let | |
| 30216 
0300b7420b07
nicer_shortest: use NameSpace.extern_flags with disabled "features" instead of internal NameSpace.get_accesses;
 wenzelm parents: 
30188diff
changeset | 352 | (* FIXME global name space!? *) | 
| 
0300b7420b07
nicer_shortest: use NameSpace.extern_flags with disabled "features" instead of internal NameSpace.get_accesses;
 wenzelm parents: 
30188diff
changeset | 353 | val space = Facts.space_of (PureThy.facts_of (ProofContext.theory_of ctxt)); | 
| 29848 | 354 | |
| 30216 
0300b7420b07
nicer_shortest: use NameSpace.extern_flags with disabled "features" instead of internal NameSpace.get_accesses;
 wenzelm parents: 
30188diff
changeset | 355 | val shorten = | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33039diff
changeset | 356 | Name_Space.extern_flags | 
| 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33039diff
changeset | 357 |         {long_names = false, short_names = false, unique_names = false} space;
 | 
| 29848 | 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 = | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 366 | xs ~~ (1 upto length xs) | 
| 35408 | 367 | |> sort (Term_Ord.fast_term_ord o pairself (Thm.prop_of o #2 o #1)) | 
| 29848 | 368 | |> rem_cdups nicer | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 369 | |> sort (int_ord o pairself #2) | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 370 | |> map #1; | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 371 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26283diff
changeset | 372 | end; | 
| 22340 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 373 | |
| 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 kleing parents: 
19502diff
changeset | 374 | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 375 | (* print_theorems *) | 
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 376 | |
| 26283 | 377 | fun all_facts_of ctxt = | 
| 33381 | 378 | let | 
| 33382 | 379 | fun visible_facts facts = | 
| 380 | Facts.dest_static [] facts | |
| 381 | |> filter_out (Facts.is_concealed facts o #1); | |
| 33381 | 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; | |
| 17972 | 387 | |
| 32738 | 388 | val limit = Unsynchronized.ref 40; | 
| 25992 
928594f50893
renamed thms_containing_limit to FindTheorems.limit;
 wenzelm parents: 
25226diff
changeset | 389 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 390 | fun find_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 391 | let | 
| 30822 | 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); | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 396 | val opt_goal' = Option.map add_prems opt_goal; | 
| 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 397 | |
| 16036 | 398 | val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 399 | val filters = map (filter_criterion ctxt opt_goal') criteria; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 400 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 401 | fun find_all facts = | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 402 | let | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 403 | val raw_matches = sorted_filter filters facts; | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 404 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 405 | val matches = | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 406 | if rem_dups | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 407 | then rem_thm_dups (nicer_shortest ctxt) raw_matches | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 408 | else raw_matches; | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 409 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 410 | val len = length matches; | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 411 | val lim = the_default (! limit) opt_limit; | 
| 34088 | 412 | 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 | 413 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 414 | val find = | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 415 | if rem_dups orelse is_none opt_limit | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 416 | then find_all | 
| 30822 | 417 | 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 | 418 | |
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 419 | in find (all_facts_of ctxt) end; | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 420 | |
| 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 | 421 | |
| 
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 | 422 | fun pretty_thm ctxt (thmref, thm) = Pretty.block | 
| 
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 | 423 | [Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1, | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31687diff
changeset | 424 | Display.pretty_thm ctxt thm]; | 
| 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 | 425 | |
| 30143 | 426 | fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = | 
| 427 | let | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 428 | val start = start_timing (); | 
| 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 429 | |
| 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 430 | val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 431 | val (foundo, thms) = find_theorems ctxt opt_goal opt_limit rem_dups raw_criteria; | 
| 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 432 | val returned = length thms; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
31042diff
changeset | 433 | |
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 434 | val tally_msg = | 
| 30822 | 435 | (case foundo of | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 436 | NONE => "displaying " ^ string_of_int returned ^ " theorems" | 
| 30822 | 437 | | SOME found => | 
| 438 | "found " ^ string_of_int found ^ " theorems" ^ | |
| 439 | (if returned < found | |
| 440 |              then " (" ^ string_of_int returned ^ " displayed)"
 | |
| 441 | else "")); | |
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 442 | |
| 31687 | 443 | val end_msg = " in " ^ Time.toString (#cpu (end_timing start)) ^ " secs"; | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 444 | in | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 445 | Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) | 
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29848diff
changeset | 446 | :: Pretty.str "" :: | 
| 28900 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 kleing parents: 
28211diff
changeset | 447 |      (if null thms then [Pretty.str ("nothing found" ^ end_msg)]
 | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 448 | else | 
| 30785 
15f64e05e703
Limit the number of results returned by auto_solves.
 Timothy Bourke parents: 
30693diff
changeset | 449 | [Pretty.str (tally_msg ^ end_msg ^ ":"), Pretty.str ""] @ | 
| 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 | 450 | map (pretty_thm ctxt) thms) | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 451 | |> Pretty.chunks |> Pretty.writeln | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 452 | end; | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 453 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 454 | |
| 32798 | 455 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 456 | (** command syntax **) | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 457 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 458 | fun find_theorems_cmd ((opt_lim, rem_dups), spec) = | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 459 | Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 30822 | 460 | let | 
| 461 | val proof_state = Toplevel.enter_proof_body state; | |
| 462 | val ctxt = Proof.context_of proof_state; | |
| 33290 | 463 | val opt_goal = try Proof.simple_goal proof_state |> Option.map #goal; | 
| 30822 | 464 | in print_theorems ctxt opt_goal opt_lim rem_dups spec 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 | 465 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 466 | local | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 467 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 468 | val criterion = | 
| 36950 | 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; | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 477 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 478 | val options = | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 479 | Scan.optional | 
| 36950 | 480 |     (Parse.$$$ "(" |--
 | 
| 481 | Parse.!!! (Scan.option Parse.nat -- Scan.optional (Parse.reserved "with_dups" >> K false) true | |
| 482 | --| Parse.$$$ ")")) (NONE, true); | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 483 | in | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 484 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 485 | val _ = | 
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 486 | Outer_Syntax.improper_command "find_theorems" "print theorems meeting specified criteria" | 
| 36950 | 487 | Keyword.diag | 
| 488 | (options -- Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion)) | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 489 | >> (Toplevel.no_timing oo find_theorems_cmd)); | 
| 16033 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 490 | |
| 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 wenzelm parents: diff
changeset | 491 | 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 | 492 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29882diff
changeset | 493 | end; |