author | haftmann |
Mon, 30 Aug 2010 17:20:33 +0200 | |
changeset 38918 | 48d62f237944 |
parent 38336 | fd53ae1d4c47 |
child 42012 | 2c3fe3cbebae |
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:
30207
diff
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 |
|
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
39 |
fun filter_const _ NONE = NONE |
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
40 |
| filter_const f (SOME (c, r)) = |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
41 |
(case f c of |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
42 |
NONE => NONE |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
43 |
| SOME i => SOME (c, Int.min (r, i))); |
30143 | 44 |
|
45 |
||
46 |
(* pretty results *) |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
47 |
|
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
48 |
fun pretty_criterion (b, c) = |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
49 |
let |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
50 |
fun prfx s = if b then s else "-" ^ s; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
51 |
in |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
52 |
(case c of |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
53 |
Strict pat => Pretty.str (prfx "strict: " ^ quote pat) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
54 |
| Loose pat => Pretty.str (prfx (quote pat)) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
55 |
| Name name => Pretty.str (prfx "name: " ^ quote name)) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
56 |
end; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
57 |
|
30143 | 58 |
fun pretty_const ctxt (nm, ty) = |
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:
33301
diff
changeset
|
60 |
val ty' = Logic.unvarifyT_global ty; |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
61 |
in |
30143 | 62 |
Pretty.block |
38336
fd53ae1d4c47
standardized pretty printing of consts (e.g. see find_theorems, print_theory);
wenzelm
parents:
38335
diff
changeset
|
63 |
[Pretty.str nm, Pretty.str " ::", Pretty.brk 1, |
30143 | 64 |
Pretty.quote (Syntax.pretty_typ ctxt ty')] |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
65 |
end; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
66 |
|
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
67 |
|
30143 | 68 |
(* find_consts *) |
69 |
||
70 |
fun find_consts ctxt raw_criteria = |
|
71 |
let |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
72 |
val start = start_timing (); |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
73 |
|
29884 | 74 |
val thy = ProofContext.theory_of ctxt; |
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:
30190
diff
changeset
|
79 |
|
30207
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
80 |
fun make_pattern crit = |
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
81 |
let |
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
82 |
val raw_T = Syntax.parse_typ ctxt crit; |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
83 |
val t = |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
84 |
Syntax.check_term |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
85 |
(ProofContext.set_mode ProofContext.mode_pattern ctxt) |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
86 |
(Term.dummy_pattern raw_T); |
30207
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
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:
30207
diff
changeset
|
94 |
val sub_size = |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
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:
31687
diff
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:
30190
diff
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:
30190
diff
changeset
|
105 |
val consts = Sign.consts_of thy; |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
106 |
val (_, consts_tab) = #constants (Consts.dest consts); |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
107 |
fun eval_entry c = |
33158 | 108 |
fold filter_const (user_visible consts :: criteria) |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
109 |
(SOME (c, low_ranking)); |
29884 | 110 |
|
30143 | 111 |
val matches = |
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
112 |
Symtab.fold (cons o eval_entry) consts_tab [] |
30143 | 113 |
|> map_filter I |
114 |
|> sort (rev_order o int_ord o pairself snd) |
|
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
115 |
|> map (apsnd fst o fst); |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
116 |
|
31687 | 117 |
val end_msg = " in " ^ Time.toString (#cpu (end_timing start)) ^ " secs"; |
29884 | 118 |
in |
38335 | 119 |
Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) :: |
120 |
Pretty.str "" :: |
|
121 |
Pretty.str |
|
122 |
(if null matches |
|
123 |
then "nothing found" ^ end_msg |
|
124 |
else "found " ^ string_of_int (length matches) ^ " constant(s)" ^ end_msg ^ ":") :: |
|
125 |
Pretty.str "" :: |
|
126 |
map (pretty_const ctxt) matches |
|
127 |
end |> Pretty.chunks |> Pretty.writeln; |
|
29884 | 128 |
|
30142
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
129 |
|
30143 | 130 |
(* command syntax *) |
30142
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
131 |
|
38334 | 132 |
local |
30142
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
133 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
134 |
val criterion = |
36950 | 135 |
Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Strict || |
136 |
Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name || |
|
137 |
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:
29895
diff
changeset
|
138 |
|
38334 | 139 |
in |
140 |
||
30142
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
141 |
val _ = |
36953
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents:
36950
diff
changeset
|
142 |
Outer_Syntax.improper_command "find_consts" "search constants by type pattern" Keyword.diag |
36950 | 143 |
(Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion) |
38334 | 144 |
>> (fn spec => Toplevel.no_timing o |
145 |
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:
29895
diff
changeset
|
146 |
|
29884 | 147 |
end; |
148 |
||
38334 | 149 |
end; |
150 |