author | haftmann |
Fri, 25 Sep 2009 09:50:31 +0200 | |
changeset 32705 | 04ce6bb14d85 |
parent 31687 | 0d2f700fe5e7 |
child 32790 | a7b92f98180b |
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 |
|
15 |
val find_consts : Proof.context -> (bool * criterion) list -> unit |
|
16 |
end; |
|
17 |
||
18 |
structure FindConsts : FIND_CONSTS = |
|
19 |
struct |
|
20 |
||
30143 | 21 |
(* search criteria *) |
22 |
||
23 |
datatype criterion = |
|
24 |
Strict of string |
|
25 |
| Loose of string |
|
26 |
| Name of string; |
|
29884 | 27 |
|
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
28 |
|
30143 | 29 |
(* matching types/consts *) |
29884 | 30 |
|
30143 | 31 |
fun matches_subtype thy typat = |
32 |
let |
|
29884 | 33 |
val p = can (fn ty => Sign.typ_match thy (typat, ty) Vartab.empty); |
34 |
||
35 |
fun fs [] = false |
|
30143 | 36 |
| fs (t :: ts) = f t orelse fs ts |
29884 | 37 |
|
38 |
and f (t as Type (_, ars)) = p t orelse fs ars |
|
39 |
| f t = p t; |
|
40 |
in f end; |
|
41 |
||
30143 | 42 |
fun check_const p (nm, (ty, _)) = |
43 |
if p (nm, ty) |
|
44 |
then SOME (Term.size_of_typ ty) |
|
45 |
else NONE; |
|
29884 | 46 |
|
30143 | 47 |
fun opt_not f (c as (_, (ty, _))) = |
48 |
if is_some (f c) |
|
49 |
then NONE else SOME (Term.size_of_typ ty); |
|
29884 | 50 |
|
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
51 |
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
|
52 |
| filter_const f (SOME (c, r)) = |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
53 |
(case f c of |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
54 |
NONE => NONE |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
55 |
| SOME i => SOME (c, Int.min (r, i))); |
30143 | 56 |
|
57 |
||
58 |
(* pretty results *) |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
59 |
|
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
60 |
fun pretty_criterion (b, c) = |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
61 |
let |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
62 |
fun prfx s = if b then s else "-" ^ s; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
63 |
in |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
64 |
(case c of |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
65 |
Strict pat => Pretty.str (prfx "strict: " ^ quote pat) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
66 |
| Loose pat => Pretty.str (prfx (quote pat)) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
67 |
| Name name => Pretty.str (prfx "name: " ^ quote name)) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
68 |
end; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
69 |
|
30143 | 70 |
fun pretty_const ctxt (nm, ty) = |
71 |
let |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
72 |
val ty' = Logic.unvarifyT ty; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
73 |
in |
30143 | 74 |
Pretty.block |
75 |
[Pretty.quote (Pretty.str nm), Pretty.fbrk, |
|
76 |
Pretty.str "::", Pretty.brk 1, |
|
77 |
Pretty.quote (Syntax.pretty_typ ctxt ty')] |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
78 |
end; |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
79 |
|
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
80 |
|
30143 | 81 |
(* find_consts *) |
82 |
||
83 |
fun find_consts ctxt raw_criteria = |
|
84 |
let |
|
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
85 |
val start = start_timing (); |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
86 |
|
29884 | 87 |
val thy = ProofContext.theory_of ctxt; |
88 |
val low_ranking = 10000; |
|
89 |
||
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
90 |
fun not_internal consts (nm, _) = |
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
91 |
if member (op =) (Consts.the_tags consts nm) Markup.property_internal |
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
92 |
then NONE else SOME low_ranking; |
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
93 |
|
30207
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
94 |
fun make_pattern crit = |
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
95 |
let |
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
96 |
val raw_T = Syntax.parse_typ ctxt crit; |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
97 |
val t = |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
98 |
Syntax.check_term |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
99 |
(ProofContext.set_mode ProofContext.mode_pattern ctxt) |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
100 |
(Term.dummy_pattern raw_T); |
30207
c56d27155041
Implement Makarius's suggestion for improved type pattern parsing.
Timothy Bourke
parents:
30206
diff
changeset
|
101 |
in Term.type_of t end; |
29884 | 102 |
|
103 |
fun make_match (Strict arg) = |
|
104 |
let val qty = make_pattern arg; in |
|
30143 | 105 |
fn (_, (ty, _)) => |
106 |
let |
|
29884 | 107 |
val tye = Sign.typ_match thy (qty, ty) Vartab.empty; |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
108 |
val sub_size = |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
109 |
Vartab.fold (fn (_, (_, t)) => fn n => Term.size_of_typ t + n) tye 0; |
29884 | 110 |
in SOME sub_size end handle MATCH => NONE |
111 |
end |
|
112 |
| make_match (Loose arg) = |
|
113 |
check_const (matches_subtype thy (make_pattern arg) o snd) |
|
114 |
| make_match (Name arg) = check_const (match_string arg o fst); |
|
115 |
||
116 |
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
|
117 |
val criteria = map make_criterion raw_criteria; |
29884 | 118 |
|
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
119 |
val consts = Sign.consts_of thy; |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
120 |
val (_, consts_tab) = #constants (Consts.dest consts); |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
121 |
fun eval_entry c = |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
122 |
fold filter_const (not_internal consts :: criteria) |
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
123 |
(SOME (c, low_ranking)); |
29884 | 124 |
|
30143 | 125 |
val matches = |
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
126 |
Symtab.fold (cons o eval_entry) consts_tab [] |
30143 | 127 |
|> map_filter I |
128 |
|> sort (rev_order o int_ord o pairself snd) |
|
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
129 |
|> map (apsnd fst o fst); |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
130 |
|
31687 | 131 |
val end_msg = " in " ^ Time.toString (#cpu (end_timing start)) ^ " secs"; |
29884 | 132 |
in |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
133 |
Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
134 |
:: Pretty.str "" |
31684
d5d830979a54
minor tuning according to Isabelle/ML conventions;
wenzelm
parents:
30207
diff
changeset
|
135 |
:: (Pretty.str o implode) |
29895
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
136 |
(if null matches |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
137 |
then ["nothing found", end_msg] |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
138 |
else ["found ", (string_of_int o length) matches, |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
139 |
" constants", end_msg, ":"]) |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
140 |
:: Pretty.str "" |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
141 |
:: map (pretty_const ctxt) matches |
0e70a29d3e02
find_consts: display the search criteria. (by Timothy Bourke)
kleing
parents:
29884
diff
changeset
|
142 |
|> Pretty.chunks |
29884 | 143 |
|> Pretty.writeln |
30206
48507466d9d2
find_consts: fold in preference to foldl; hide internal constants; remove redundant exception catch
Timothy Bourke
parents:
30190
diff
changeset
|
144 |
end; |
29884 | 145 |
|
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 |
|
30143 | 147 |
(* 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
|
148 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
149 |
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:
29895
diff
changeset
|
150 |
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:
29895
diff
changeset
|
151 |
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:
29895
diff
changeset
|
152 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
153 |
local |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
154 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
155 |
structure P = OuterParse and K = OuterKeyword; |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
156 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
157 |
val criterion = |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
158 |
P.reserved "strict" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Strict || |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
159 |
P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Name || |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
160 |
P.xname >> Loose; |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
161 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
162 |
in |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
163 |
|
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
164 |
val _ = |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
165 |
OuterSyntax.improper_command "find_consts" "search constants by type pattern" K.diag |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
166 |
(Scan.repeat (((Scan.option P.minus >> is_none) -- criterion)) |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
167 |
>> (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:
29895
diff
changeset
|
168 |
|
29884 | 169 |
end; |
170 |
||
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
|
171 |
end; |
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents:
29895
diff
changeset
|
172 |