| author | wenzelm | 
| Fri, 07 Mar 2014 20:46:27 +0100 | |
| changeset 55987 | 52c22561996d | 
| parent 52703 | d68fd63bf082 | 
| child 56025 | d74fed45fa8b | 
| 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 | |
| 29884 | 14 | val find_consts : Proof.context -> (bool * criterion) list -> unit | 
| 15 | end; | |
| 16 | ||
| 33301 | 17 | structure Find_Consts : FIND_CONSTS = | 
| 29884 | 18 | struct | 
| 19 | ||
| 30143 | 20 | (* search criteria *) | 
| 21 | ||
| 22 | datatype criterion = | |
| 23 | Strict of string | |
| 24 | | Loose of string | |
| 25 | | Name of string; | |
| 29884 | 26 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 27 | |
| 30143 | 28 | (* matching types/consts *) | 
| 29884 | 29 | |
| 30143 | 30 | fun matches_subtype thy typat = | 
| 38335 | 31 | Term.exists_subtype (fn ty => Sign.typ_instance thy (ty, typat)); | 
| 29884 | 32 | |
| 38335 | 33 | fun check_const pred (nm, (ty, _)) = | 
| 34 | if pred (nm, ty) then SOME (Term.size_of_typ ty) else NONE; | |
| 29884 | 35 | |
| 30143 | 36 | fun opt_not f (c as (_, (ty, _))) = | 
| 38335 | 37 | if is_some (f c) then NONE else SOME (Term.size_of_typ ty); | 
| 29884 | 38 | |
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 39 | fun filter_const _ _ NONE = NONE | 
| 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 40 | | filter_const c f (SOME rank) = | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 41 | (case f c of | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 42 | NONE => NONE | 
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 43 | | SOME i => SOME (Int.min (rank, i))); | 
| 30143 | 44 | |
| 45 | ||
| 46 | (* pretty results *) | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 47 | |
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 48 | fun pretty_criterion (b, c) = | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 49 | let | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 50 | fun prfx s = if b then s else "-" ^ s; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 51 | in | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 52 | (case c of | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 53 | Strict pat => Pretty.str (prfx "strict: " ^ quote pat) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 54 | | Loose pat => Pretty.str (prfx (quote pat)) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 55 | | Name name => Pretty.str (prfx "name: " ^ quote name)) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 56 | end; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 57 | |
| 49886 | 58 | fun pretty_const ctxt (c, ty) = | 
| 30143 | 59 | 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 | 60 | val ty' = Logic.unvarifyT_global ty; | 
| 49886 | 61 | val consts_space = Consts.space_of (Sign.consts_of (Proof_Context.theory_of ctxt)); | 
| 62 | val markup = Name_Space.markup consts_space c; | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 63 | in | 
| 30143 | 64 | Pretty.block | 
| 49886 | 65 | [Pretty.mark markup (Pretty.str c), Pretty.str " ::", Pretty.brk 1, | 
| 30143 | 66 | Pretty.quote (Syntax.pretty_typ ctxt ty')] | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 67 | end; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 68 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 69 | |
| 30143 | 70 | (* find_consts *) | 
| 71 | ||
| 72 | fun find_consts ctxt raw_criteria = | |
| 73 | let | |
| 42360 | 74 | val thy = Proof_Context.theory_of ctxt; | 
| 29884 | 75 | val low_ranking = 10000; | 
| 76 | ||
| 33158 | 77 | fun user_visible consts (nm, _) = | 
| 78 | if Consts.is_concealed consts nm then NONE else SOME low_ranking; | |
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 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 | |
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 105 | val consts = Sign.consts_of thy; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 106 | val (_, consts_tab) = #constants (Consts.dest consts); | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 107 | fun eval_entry c = | 
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 108 | fold (filter_const c) (user_visible consts :: criteria) (SOME low_ranking); | 
| 29884 | 109 | |
| 30143 | 110 | val matches = | 
| 46979 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 111 | Symtab.fold (fn c => (case eval_entry c of NONE => I | SOME rank => cons (rank, c))) | 
| 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 112 | consts_tab [] | 
| 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 113 | |> sort (prod_ord (rev_order o int_ord) (string_ord o pairself fst)) | 
| 
ef4b0d6b2fb6
sort via string_ord (as secondary key), not fast_string_ord via Symtab.fold;
 wenzelm parents: 
46961diff
changeset | 114 | |> map (apsnd fst o snd); | 
| 29884 | 115 | in | 
| 38335 | 116 | Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) :: | 
| 117 | Pretty.str "" :: | |
| 118 | Pretty.str | |
| 119 | (if null matches | |
| 46716 
c45a4427db39
discontinued slightly odd built-in timing (cf. 53fd5cc685b4) -- the Isar toplevel does that already (e.g. via Toplevel.timing or Toplevel.profiling);
 wenzelm parents: 
46713diff
changeset | 120 | then "nothing found" | 
| 
c45a4427db39
discontinued slightly odd built-in timing (cf. 53fd5cc685b4) -- the Isar toplevel does that already (e.g. via Toplevel.timing or Toplevel.profiling);
 wenzelm parents: 
46713diff
changeset | 121 | else "found " ^ string_of_int (length matches) ^ " constant(s):") :: | 
| 38335 | 122 | Pretty.str "" :: | 
| 123 | map (pretty_const ctxt) matches | |
| 52703 
d68fd63bf082
tuned messages -- avoid text folds stemming from Pretty.chunks;
 wenzelm parents: 
51658diff
changeset | 124 | end |> Pretty.fbreaks |> curry Pretty.blk 0 |> Pretty.writeln; | 
| 29884 | 125 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 126 | |
| 30143 | 127 | (* 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 | 128 | |
| 38334 | 129 | 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 | 130 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 131 | val criterion = | 
| 36950 | 132 | Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Strict || | 
| 133 | Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name || | |
| 134 | Parse.xname >> 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 | 135 | |
| 38334 | 136 | in | 
| 137 | ||
| 30142 
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 _ = | 
| 50214 | 139 |   Outer_Syntax.improper_command @{command_spec "find_consts"}
 | 
| 140 | "find constants by name/type patterns" | |
| 36950 | 141 | (Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion) | 
| 51658 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 wenzelm parents: 
50214diff
changeset | 142 | >> (fn spec => Toplevel.keep (fn state => find_consts (Toplevel.context_of state) spec))); | 
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 143 | |
| 29884 | 144 | end; | 
| 145 | ||
| 38334 | 146 | end; | 
| 147 |