| author | wenzelm | 
| Sat, 15 Oct 2016 22:45:27 +0200 | |
| changeset 64236 | 358f9ff08681 | 
| parent 63429 | baedd4724f08 | 
| child 64556 | 851ae0e7b09c | 
| permissions | -rw-r--r-- | 
| 30143 | 1 | (* Title: Pure/Tools/find_consts.ML | 
| 29884 | 2 | Author: Timothy Bourke and Gerwin Klein, NICTA | 
| 3 | ||
| 30143 | 4 | Hoogle-like (http://www-users.cs.york.ac.uk/~ndm/hoogle) searching by | 
| 5 | type over constants, but matching is not fuzzy. | |
| 29884 | 6 | *) | 
| 7 | ||
| 8 | signature FIND_CONSTS = | |
| 9 | sig | |
| 30143 | 10 | datatype criterion = | 
| 11 | Strict of string | |
| 12 | | Loose of string | |
| 13 | | Name of string | |
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61335diff
changeset | 14 | val pretty_consts: Proof.context -> (bool * criterion) list -> Pretty.T | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61335diff
changeset | 15 | val query_parser: (bool * criterion) list parser | 
| 56758 | 16 | val read_query: Position.T -> string -> (bool * criterion) list | 
| 29884 | 17 | val find_consts : Proof.context -> (bool * criterion) list -> unit | 
| 18 | end; | |
| 19 | ||
| 33301 | 20 | structure Find_Consts : FIND_CONSTS = | 
| 29884 | 21 | struct | 
| 22 | ||
| 30143 | 23 | (* search criteria *) | 
| 24 | ||
| 25 | datatype criterion = | |
| 26 | Strict of string | |
| 27 | | Loose of string | |
| 28 | | Name of string; | |
| 29884 | 29 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 30 | |
| 30143 | 31 | (* matching types/consts *) | 
| 29884 | 32 | |
| 30143 | 33 | fun matches_subtype thy typat = | 
| 38335 | 34 | Term.exists_subtype (fn ty => Sign.typ_instance thy (ty, typat)); | 
| 29884 | 35 | |
| 61335 | 36 | fun check_const pred (c, (ty, _)) = | 
| 37 | if pred (c, ty) then SOME (Term.size_of_typ ty) else NONE; | |
| 29884 | 38 | |
| 30143 | 39 | fun opt_not f (c as (_, (ty, _))) = | 
| 38335 | 40 | if is_some (f c) then NONE else SOME (Term.size_of_typ ty); | 
| 29884 | 41 | |
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 42 | fun filter_const _ _ NONE = NONE | 
| 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 43 | | filter_const c f (SOME rank) = | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 44 | (case f c of | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 45 | NONE => NONE | 
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 46 | | SOME i => SOME (Int.min (rank, i))); | 
| 30143 | 47 | |
| 48 | ||
| 49 | (* pretty results *) | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 50 | |
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 51 | fun pretty_criterion (b, c) = | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 52 | let | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 53 | fun prfx s = if b then s else "-" ^ s; | 
| 60667 
d86c449d30ba
plain string output, without funny control chars;
 wenzelm parents: 
60664diff
changeset | 54 | val show_pat = quote o Input.source_content o Syntax.read_input; | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 55 | in | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 56 | (case c of | 
| 60667 
d86c449d30ba
plain string output, without funny control chars;
 wenzelm parents: 
60664diff
changeset | 57 | Strict pat => Pretty.str (prfx "strict: " ^ show_pat pat) | 
| 
d86c449d30ba
plain string output, without funny control chars;
 wenzelm parents: 
60664diff
changeset | 58 | | Loose pat => Pretty.str (prfx (show_pat pat)) | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 59 | | Name name => Pretty.str (prfx "name: " ^ quote name)) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 60 | end; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 61 | |
| 49886 | 62 | fun pretty_const ctxt (c, ty) = | 
| 30143 | 63 | let | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
33301diff
changeset | 64 | val ty' = Logic.unvarifyT_global ty; | 
| 61335 | 65 | val markup = Name_Space.markup (Proof_Context.const_space ctxt) c; | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 66 | in | 
| 30143 | 67 | Pretty.block | 
| 49886 | 68 | [Pretty.mark markup (Pretty.str c), Pretty.str " ::", Pretty.brk 1, | 
| 30143 | 69 | Pretty.quote (Syntax.pretty_typ ctxt ty')] | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 70 | end; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 71 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 72 | |
| 30143 | 73 | (* find_consts *) | 
| 74 | ||
| 56758 | 75 | fun pretty_consts ctxt raw_criteria = | 
| 30143 | 76 | let | 
| 42360 | 77 | val thy = Proof_Context.theory_of ctxt; | 
| 29884 | 78 | val low_ranking = 10000; | 
| 79 | ||
| 30207 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 80 | fun make_pattern crit = | 
| 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 81 | let | 
| 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 82 | val raw_T = Syntax.parse_typ ctxt crit; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 83 | val t = | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 84 | Syntax.check_term | 
| 42360 | 85 | (Proof_Context.set_mode Proof_Context.mode_pattern ctxt) | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 86 | (Term.dummy_pattern raw_T); | 
| 30207 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 87 | in Term.type_of t end; | 
| 29884 | 88 | |
| 89 | fun make_match (Strict arg) = | |
| 90 | let val qty = make_pattern arg; in | |
| 30143 | 91 | fn (_, (ty, _)) => | 
| 92 | let | |
| 29884 | 93 | val tye = Sign.typ_match thy (qty, ty) Vartab.empty; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 94 | val sub_size = | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 95 | Vartab.fold (fn (_, (_, t)) => fn n => Term.size_of_typ t + n) tye 0; | 
| 32790 
a7b92f98180b
handle Type.TYPE_MATCH, not arbitrary exceptions;
 wenzelm parents: 
31687diff
changeset | 96 | in SOME sub_size end handle Type.TYPE_MATCH => NONE | 
| 29884 | 97 | end | 
| 98 | | make_match (Loose arg) = | |
| 99 | check_const (matches_subtype thy (make_pattern arg) o snd) | |
| 100 | | make_match (Name arg) = check_const (match_string arg o fst); | |
| 101 | ||
| 102 | fun make_criterion (b, crit) = (if b then I else opt_not) (make_match crit); | |
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 103 | val criteria = map make_criterion raw_criteria; | 
| 29884 | 104 | |
| 61335 | 105 | fun user_visible (c, _) = | 
| 106 | if Consts.is_concealed (Proof_Context.consts_of ctxt) c | |
| 107 | then NONE else SOME low_ranking; | |
| 108 | ||
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 109 | fun eval_entry c = | 
| 61335 | 110 | fold (filter_const c) (user_visible :: criteria) (SOME low_ranking); | 
| 29884 | 111 | |
| 61335 | 112 |     val {constants, ...} = Consts.dest (Sign.consts_of thy);
 | 
| 30143 | 113 | val matches = | 
| 56062 | 114 | fold (fn c => (case eval_entry c of NONE => I | SOME rank => cons (rank, c))) constants [] | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 115 | |> sort (prod_ord (rev_order o int_ord) (string_ord o apply2 fst)) | 
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 116 | |> map (apsnd fst o snd); | 
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 117 | |
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 118 | val position_markup = Position.markup (Position.thread_data ()) Markup.position; | 
| 29884 | 119 | in | 
| 56891 
48899c43b07d
tuned message -- more context for detached window etc.;
 wenzelm parents: 
56763diff
changeset | 120 | Pretty.block | 
| 56912 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 121 | (Pretty.fbreaks | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 122 | (Pretty.mark position_markup (Pretty.keyword1 "find_consts") :: | 
| 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56908diff
changeset | 123 | map pretty_criterion raw_criteria)) :: | 
| 38335 | 124 | Pretty.str "" :: | 
| 56908 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 125 | (if null matches then [Pretty.str "found nothing"] | 
| 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 126 | else | 
| 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 127 |        Pretty.str ("found " ^ string_of_int (length matches) ^ " constant(s):") ::
 | 
| 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 wenzelm parents: 
56891diff
changeset | 128 | grouped 10 Par_List.map (Pretty.item o single o pretty_const ctxt) matches) | 
| 61335 | 129 | end |> Pretty.chunks; | 
| 56758 | 130 | |
| 131 | fun find_consts ctxt args = Pretty.writeln (pretty_consts ctxt args); | |
| 29884 | 132 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 133 | |
| 30143 | 134 | (* command syntax *) | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 135 | |
| 38334 | 136 | local | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 137 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 138 | val criterion = | 
| 62969 | 139 | Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.name) >> Name || | 
| 60664 
263a8f15faf3
proper outer syntax category, e.g. relevant for PIDE markup;
 wenzelm parents: 
60610diff
changeset | 140 | Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.typ) >> Strict || | 
| 
263a8f15faf3
proper outer syntax category, e.g. relevant for PIDE markup;
 wenzelm parents: 
60610diff
changeset | 141 | Parse.typ >> Loose; | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 142 | |
| 63429 | 143 | val query_keywords = | 
| 144 |   Keyword.add_keywords [((":", @{here}), Keyword.no_spec)] Keyword.empty_keywords;
 | |
| 58905 | 145 | |
| 38334 | 146 | in | 
| 147 | ||
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61335diff
changeset | 148 | val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); | 
| 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61335diff
changeset | 149 | |
| 56758 | 150 | fun read_query pos str = | 
| 59083 | 151 | Token.explode query_keywords pos str | 
| 56758 | 152 | |> filter Token.is_proper | 
| 62848 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 wenzelm parents: 
61335diff
changeset | 153 | |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof))) | 
| 56758 | 154 | |> #1; | 
| 155 | ||
| 29884 | 156 | end; | 
| 157 | ||
| 56758 | 158 | |
| 159 | (* PIDE query operation *) | |
| 160 | ||
| 161 | 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: 
59936diff
changeset | 162 |   Query_Operation.register {name = "find_consts", pri = Task_Queue.urgent_pri}
 | 
| 61223 | 163 |     (fn {state, 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: 
59936diff
changeset | 164 | (case try Toplevel.context_of state of | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59936diff
changeset | 165 | SOME ctxt => | 
| 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
59936diff
changeset | 166 | 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: 
59936diff
changeset | 167 | val [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: 
59936diff
changeset | 168 | val query = read_query Position.none query_arg; | 
| 61223 | 169 | in writeln_result (Pretty.string_of (pretty_consts ctxt query)) 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: 
59936diff
changeset | 170 | | NONE => error "Unknown context")); | 
| 56758 | 171 | |
| 38334 | 172 | end; |