| author | haftmann | 
| Tue, 10 Aug 2010 14:53:41 +0200 | |
| changeset 38312 | 9dd57db3c0f2 | 
| parent 36953 | 2af1ad9aa1a3 | 
| child 38334 | c677c2c1d333 | 
| 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 = | 
| 31 | let | |
| 29884 | 32 | val p = can (fn ty => Sign.typ_match thy (typat, ty) Vartab.empty); | 
| 33 | ||
| 34 | fun fs [] = false | |
| 30143 | 35 | | fs (t :: ts) = f t orelse fs ts | 
| 29884 | 36 | |
| 37 | and f (t as Type (_, ars)) = p t orelse fs ars | |
| 38 | | f t = p t; | |
| 39 | in f end; | |
| 40 | ||
| 30143 | 41 | fun check_const p (nm, (ty, _)) = | 
| 42 | if p (nm, ty) | |
| 43 | then SOME (Term.size_of_typ ty) | |
| 44 | else NONE; | |
| 29884 | 45 | |
| 30143 | 46 | fun opt_not f (c as (_, (ty, _))) = | 
| 47 | if is_some (f c) | |
| 48 | then NONE else SOME (Term.size_of_typ ty); | |
| 29884 | 49 | |
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 50 | fun filter_const _ NONE = NONE | 
| 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 51 | | filter_const f (SOME (c, r)) = | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 52 | (case f c of | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 53 | NONE => NONE | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 54 | | SOME i => SOME (c, Int.min (r, i))); | 
| 30143 | 55 | |
| 56 | ||
| 57 | (* pretty results *) | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 58 | |
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 59 | fun pretty_criterion (b, c) = | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 60 | let | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 61 | fun prfx s = if b then s else "-" ^ s; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 62 | in | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 63 | (case c of | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 64 | Strict pat => Pretty.str (prfx "strict: " ^ quote pat) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 65 | | Loose pat => Pretty.str (prfx (quote pat)) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 66 | | Name name => Pretty.str (prfx "name: " ^ quote name)) | 
| 
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 | |
| 30143 | 69 | fun pretty_const ctxt (nm, ty) = | 
| 70 | 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 | 71 | val ty' = Logic.unvarifyT_global ty; | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 72 | in | 
| 30143 | 73 | Pretty.block | 
| 74 | [Pretty.quote (Pretty.str nm), Pretty.fbrk, | |
| 75 | Pretty.str "::", Pretty.brk 1, | |
| 76 | Pretty.quote (Syntax.pretty_typ ctxt ty')] | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 77 | end; | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 78 | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 79 | |
| 30143 | 80 | (* find_consts *) | 
| 81 | ||
| 82 | fun find_consts ctxt raw_criteria = | |
| 83 | let | |
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 84 | val start = start_timing (); | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 85 | |
| 29884 | 86 | val thy = ProofContext.theory_of ctxt; | 
| 87 | val low_ranking = 10000; | |
| 88 | ||
| 33158 | 89 | fun user_visible consts (nm, _) = | 
| 90 | 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 | 91 | |
| 30207 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 92 | fun make_pattern crit = | 
| 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 93 | let | 
| 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 94 | val raw_T = Syntax.parse_typ ctxt crit; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 95 | val t = | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 96 | Syntax.check_term | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 97 | (ProofContext.set_mode ProofContext.mode_pattern ctxt) | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 98 | (Term.dummy_pattern raw_T); | 
| 30207 
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
 Timothy Bourke parents: 
30206diff
changeset | 99 | in Term.type_of t end; | 
| 29884 | 100 | |
| 101 | fun make_match (Strict arg) = | |
| 102 | let val qty = make_pattern arg; in | |
| 30143 | 103 | fn (_, (ty, _)) => | 
| 104 | let | |
| 29884 | 105 | val tye = Sign.typ_match thy (qty, ty) Vartab.empty; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 106 | val sub_size = | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 107 | 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 | 108 | in SOME sub_size end handle Type.TYPE_MATCH => NONE | 
| 29884 | 109 | end | 
| 110 | | make_match (Loose arg) = | |
| 111 | check_const (matches_subtype thy (make_pattern arg) o snd) | |
| 112 | | make_match (Name arg) = check_const (match_string arg o fst); | |
| 113 | ||
| 114 | 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 | 115 | val criteria = map make_criterion raw_criteria; | 
| 29884 | 116 | |
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 117 | val consts = Sign.consts_of thy; | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 118 | val (_, consts_tab) = #constants (Consts.dest consts); | 
| 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 119 | fun eval_entry c = | 
| 33158 | 120 | fold filter_const (user_visible consts :: criteria) | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 121 | (SOME (c, low_ranking)); | 
| 29884 | 122 | |
| 30143 | 123 | val matches = | 
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 124 | Symtab.fold (cons o eval_entry) consts_tab [] | 
| 30143 | 125 | |> map_filter I | 
| 126 | |> sort (rev_order o int_ord o pairself snd) | |
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 127 | |> map (apsnd fst o fst); | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 128 | |
| 31687 | 129 | val end_msg = " in " ^ Time.toString (#cpu (end_timing start)) ^ " secs"; | 
| 29884 | 130 | in | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 131 | Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 132 | :: Pretty.str "" | 
| 31684 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 wenzelm parents: 
30207diff
changeset | 133 | :: (Pretty.str o implode) | 
| 29895 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 134 | (if null matches | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 135 | then ["nothing found", end_msg] | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 136 | else ["found ", (string_of_int o length) matches, | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 137 | " constants", end_msg, ":"]) | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 138 | :: Pretty.str "" | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 139 | :: map (pretty_const ctxt) matches | 
| 
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
 kleing parents: 
29884diff
changeset | 140 | |> Pretty.chunks | 
| 29884 | 141 | |> Pretty.writeln | 
| 30206 
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
 Timothy Bourke parents: 
30190diff
changeset | 142 | end; | 
| 29884 | 143 | |
| 30142 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 144 | |
| 30143 | 145 | (* 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 | 146 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 147 | fun find_consts_cmd spec = | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 148 | Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 149 | find_consts (Proof.context_of (Toplevel.enter_proof_body state)) spec); | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 150 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 151 | val criterion = | 
| 36950 | 152 | Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Strict || | 
| 153 | Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name || | |
| 154 | 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 | 155 | |
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 156 | val _ = | 
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 157 | Outer_Syntax.improper_command "find_consts" "search constants by type pattern" Keyword.diag | 
| 36950 | 158 | (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: 
29895diff
changeset | 159 | >> (Toplevel.no_timing oo find_consts_cmd)); | 
| 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 wenzelm parents: 
29895diff
changeset | 160 | |
| 29884 | 161 | end; | 
| 162 |