| author | wenzelm | 
| Sat, 10 Aug 2024 21:13:37 +0200 | |
| changeset 80687 | 9b29c5d7aae4 | 
| parent 80331 | 6f25a035069c | 
| child 80873 | e71cb37c7395 | 
| permissions | -rw-r--r-- | 
| 30143 | 1  | 
(* Title: Pure/Tools/find_theorems.ML  | 
| 26283 | 2  | 
Author: Rafal Kolanski and Gerwin Klein, NICTA  | 
| 46718 | 3  | 
Author: Lars Noschinski and Alexander Krauss, TU Muenchen  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
Retrieve theorems from proof context.  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
signature FIND_THEOREMS =  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 16036 | 10  | 
datatype 'term criterion =  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
11  | 
Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term  | 
| 43070 | 12  | 
  type 'term query = {
 | 
13  | 
goal: thm option,  | 
|
14  | 
limit: int option,  | 
|
15  | 
rem_dups: bool,  | 
|
16  | 
criteria: (bool * 'term criterion) list  | 
|
17  | 
}  | 
|
| 
62848
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
18  | 
val query_parser: (bool * string criterion) list parser  | 
| 52925 | 19  | 
val read_query: Position.T -> string -> (bool * string criterion) list  | 
| 80331 | 20  | 
val all_facts_of: Proof.context -> (Thm_Name.T * thm) list  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
21  | 
val find_theorems: Proof.context -> thm option -> int option -> bool ->  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
22  | 
(bool * term criterion) list -> int option * (Thm_Name.T * thm) list  | 
| 43067 | 23  | 
val find_theorems_cmd: Proof.context -> thm option -> int option -> bool ->  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
24  | 
(bool * string criterion) list -> int option * (Thm_Name.T * thm) list  | 
| 
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
25  | 
val pretty_thm: Proof.context -> Thm_Name.T * thm -> Pretty.T  | 
| 
62848
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
26  | 
val pretty_theorems: Proof.state ->  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
27  | 
int option -> bool -> (bool * string criterion) list -> Pretty.T  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
28  | 
val proof_state: Toplevel.state -> Proof.state  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
|
| 33301 | 31  | 
structure Find_Theorems: FIND_THEOREMS =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
struct  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
(** search criteria **)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
|
| 16036 | 36  | 
datatype 'term criterion =  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
37  | 
Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term;  | 
| 16036 | 38  | 
|
39  | 
fun read_criterion _ (Name name) = Name name  | 
|
40  | 
| read_criterion _ Intro = Intro  | 
|
41  | 
| read_criterion _ Elim = Elim  | 
|
42  | 
| read_criterion _ Dest = Dest  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
43  | 
| read_criterion _ Solves = Solves  | 
| 42360 | 44  | 
| read_criterion ctxt (Simp str) = Simp (Proof_Context.read_term_pattern ctxt str)  | 
| 
59097
 
4b05ce4c84b0
revert "better" handling of abbreviation from c61fe520602b
 
haftmann 
parents: 
59096 
diff
changeset
 | 
45  | 
| read_criterion ctxt (Pattern str) = Pattern (Proof_Context.read_term_pattern ctxt str);  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
|
| 16036 | 47  | 
fun pretty_criterion ctxt (b, c) =  | 
48  | 
let  | 
|
49  | 
fun prfx s = if b then s else "-" ^ s;  | 
|
50  | 
in  | 
|
51  | 
(case c of  | 
|
52  | 
Name name => Pretty.str (prfx "name: " ^ quote name)  | 
|
53  | 
| Intro => Pretty.str (prfx "intro")  | 
|
54  | 
| Elim => Pretty.str (prfx "elim")  | 
|
55  | 
| Dest => Pretty.str (prfx "dest")  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
56  | 
| Solves => Pretty.str (prfx "solves")  | 
| 16088 | 57  | 
| Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,  | 
| 24920 | 58  | 
Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]  | 
| 56914 | 59  | 
| Pattern pat => Pretty.enclose (prfx "\"") "\""  | 
| 24920 | 60  | 
[Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])  | 
| 16036 | 61  | 
end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
63  | 
|
| 43620 | 64  | 
|
| 43070 | 65  | 
(** queries **)  | 
66  | 
||
67  | 
type 'term query = {
 | 
|
68  | 
goal: thm option,  | 
|
69  | 
limit: int option,  | 
|
70  | 
rem_dups: bool,  | 
|
71  | 
criteria: (bool * 'term criterion) list  | 
|
72  | 
};  | 
|
73  | 
||
74  | 
fun map_criteria f {goal, limit, rem_dups, criteria} =
 | 
|
| 46718 | 75  | 
  {goal = goal, limit = limit, rem_dups = rem_dups, criteria = f criteria};
 | 
| 43070 | 76  | 
|
| 43620 | 77  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
(** search criterion filters **)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
80  | 
(*generated filters are to be of the form  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
81  | 
input: (Thm_Name.T * thm)  | 
| 53632 | 82  | 
output: (p:int, s:int, t:int) option, where  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
83  | 
NONE indicates no match  | 
| 17106 | 84  | 
p is the primary sorting criterion  | 
| 53632 | 85  | 
(eg. size of term)  | 
86  | 
s is the secondary sorting criterion  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
87  | 
(eg. number of assumptions in the theorem)  | 
| 53632 | 88  | 
t is the tertiary sorting criterion  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
89  | 
(eg. size of the substitution for intro, elim and dest)  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
90  | 
when applying a set of filters to a thm, fold results in:  | 
| 53632 | 91  | 
(max p, max s, sum of all t)  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
92  | 
*)  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
93  | 
|
| 16088 | 94  | 
|
95  | 
(* matching theorems *)  | 
|
| 17106 | 96  | 
|
| 59970 | 97  | 
fun is_nontrivial ctxt = Term.is_Const o Term.head_of o Object_Logic.drop_judgment ctxt;  | 
| 16088 | 98  | 
|
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
99  | 
(*extract terms from term_src, refine them to the parts that concern us,  | 
| 
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
100  | 
if po try match them against obj else vice versa.  | 
| 
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
101  | 
trivial matches are ignored.  | 
| 
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
102  | 
returns: smallest substitution size*)  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
103  | 
fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =  | 
| 16088 | 104  | 
let  | 
| 42360 | 105  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 16088 | 106  | 
|
| 16486 | 107  | 
fun matches pat =  | 
| 59970 | 108  | 
is_nontrivial ctxt pat andalso  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
109  | 
Pattern.matches thy (if po then (pat, obj) else (obj, pat));  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
110  | 
|
| 52940 | 111  | 
fun subst_size pat =  | 
| 18184 | 112  | 
let val (_, subst) =  | 
113  | 
Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)  | 
|
| 17205 | 114  | 
in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;  | 
| 16088 | 115  | 
|
| 52941 | 116  | 
fun best_match [] = NONE  | 
117  | 
| best_match xs = SOME (foldl1 Int.min xs);  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
118  | 
|
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
119  | 
val match_thm = matches o refine_term;  | 
| 16486 | 120  | 
in  | 
| 52940 | 121  | 
map (subst_size o refine_term) (filter match_thm (extract_terms term_src))  | 
| 52941 | 122  | 
|> best_match  | 
| 16088 | 123  | 
end;  | 
124  | 
||
125  | 
||
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
(* filter_name *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
|
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
128  | 
fun filter_name str_pat (thm_name: Thm_Name.T, _) =  | 
| 
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
129  | 
if match_string str_pat (#1 thm_name)  | 
| 53632 | 130  | 
then SOME (0, 0, 0) else NONE;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
132  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
133  | 
(* filter intro/elim/dest/solves rules *)  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
135  | 
fun filter_dest ctxt goal (_, thm) =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
let  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
137  | 
val extract_dest =  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
138  | 
(fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
hd o Logic.strip_imp_prems);  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
val prems = Logic.prems_of_goal goal 1;  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
141  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
142  | 
fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19476 
diff
changeset
 | 
143  | 
val successful = prems |> map_filter try_subst;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
in  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
145  | 
(*if possible, keep best substitution (one with smallest size)*)  | 
| 17106 | 146  | 
(*dest rules always have assumptions, so a dest with one  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
147  | 
assumption is as good as an intro rule with none*)  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
148  | 
if not (null successful) then  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
149  | 
SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
150  | 
else NONE  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
153  | 
fun filter_intro ctxt goal (_, thm) =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
let  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
155  | 
val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);  | 
| 16036 | 156  | 
val concl = Logic.concl_of_goal goal 1;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
in  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
158  | 
(case is_matching_thm extract_intro ctxt true concl thm of  | 
| 59096 | 159  | 
SOME k => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, k)  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
160  | 
| NONE => NONE)  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
163  | 
fun filter_elim ctxt goal (_, thm) =  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
164  | 
if Thm.nprems_of thm > 0 then  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
165  | 
let  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
166  | 
val rule = Thm.full_prop_of thm;  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
167  | 
val prems = Logic.prems_of_goal goal 1;  | 
| 
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
168  | 
val goal_concl = Logic.concl_of_goal goal 1;  | 
| 26283 | 169  | 
val rule_mp = hd (Logic.strip_imp_prems rule);  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
170  | 
val rule_concl = Logic.strip_imp_concl rule;  | 
| 57690 | 171  | 
      fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);  (* FIXME ?!? *)
 | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
172  | 
val rule_tree = combine rule_mp rule_concl;  | 
| 26283 | 173  | 
fun goal_tree prem = combine prem goal_concl;  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
174  | 
fun try_subst prem = is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19476 
diff
changeset
 | 
175  | 
val successful = prems |> map_filter try_subst;  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
176  | 
in  | 
| 32798 | 177  | 
(*elim rules always have assumptions, so an elim with one  | 
178  | 
assumption is as good as an intro rule with none*)  | 
|
| 59970 | 179  | 
if is_nontrivial ctxt (Thm.major_prem_of thm) andalso not (null successful) then  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
180  | 
SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
181  | 
else NONE  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
182  | 
end  | 
| 46718 | 183  | 
else NONE;  | 
| 16036 | 184  | 
|
| 30143 | 185  | 
fun filter_solves ctxt goal =  | 
186  | 
let  | 
|
| 
64983
 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 
wenzelm 
parents: 
64556 
diff
changeset
 | 
187  | 
val thy' = Proof_Context.theory_of ctxt  | 
| 
 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 
wenzelm 
parents: 
64556 
diff
changeset
 | 
188  | 
|> Context_Position.set_visible_global false;  | 
| 
 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 
wenzelm 
parents: 
64556 
diff
changeset
 | 
189  | 
val ctxt' = Proof_Context.transfer thy' ctxt  | 
| 
 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 
wenzelm 
parents: 
64556 
diff
changeset
 | 
190  | 
|> Context_Position.set_visible false;  | 
| 
52704
 
b824497c8e86
modify background theory where it is actually required (cf. 51dfdcd88e84);
 
wenzelm 
parents: 
52703 
diff
changeset
 | 
191  | 
val goal' = Thm.transfer thy' goal;  | 
| 
 
b824497c8e86
modify background theory where it is actually required (cf. 51dfdcd88e84);
 
wenzelm 
parents: 
52703 
diff
changeset
 | 
192  | 
|
| 52941 | 193  | 
fun limited_etac thm i =  | 
| 67147 | 194  | 
Seq.take (Options.default_int \<^system_option>\<open>find_theorems_tactic_limit\<close>) o  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59098 
diff
changeset
 | 
195  | 
eresolve_tac ctxt' [thm] i;  | 
| 30143 | 196  | 
fun try_thm thm =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59098 
diff
changeset
 | 
197  | 
if Thm.no_prems thm then resolve_tac ctxt' [thm] 1 goal'  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53633 
diff
changeset
 | 
198  | 
else  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53633 
diff
changeset
 | 
199  | 
(limited_etac thm THEN_ALL_NEW (Goal.norm_hhf_tac ctxt' THEN' Method.assm_tac ctxt'))  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53633 
diff
changeset
 | 
200  | 
1 goal';  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
201  | 
in  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
202  | 
fn (_, thm) =>  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
203  | 
if is_some (Seq.pull (try_thm thm))  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
204  | 
then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
205  | 
else NONE  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
206  | 
end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
208  | 
|
| 16074 | 209  | 
(* filter_simp *)  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
211  | 
fun filter_simp ctxt t (_, thm) =  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
212  | 
let  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
213  | 
val mksimps = Simplifier.mksimps ctxt;  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
214  | 
val extract_simp =  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
215  | 
(map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl);  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
216  | 
in  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
217  | 
(case is_matching_thm extract_simp ctxt false t thm of  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
218  | 
SOME ss => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, ss)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
219  | 
| NONE => NONE)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
220  | 
end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
(* filter_pattern *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
|
| 
59098
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
225  | 
fun expand_abs t =  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
226  | 
let  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
227  | 
val m = Term.maxidx_of_term t + 1;  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
228  | 
val vs = strip_abs_vars t;  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
229  | 
val ts = map_index (fn (k, (_, T)) => Var ((Name.aT, m + k), T)) vs;  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
230  | 
in betapplys (t, ts) end;  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
231  | 
|
| 32798 | 232  | 
fun get_names t = Term.add_const_names t (Term.add_free_names t []);  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
233  | 
|
| 
59095
 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 
haftmann 
parents: 
59083 
diff
changeset
 | 
234  | 
(* Does pat match a subterm of obj? *)  | 
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
235  | 
fun matches_subterm ctxt (pat, obj) =  | 
| 
59095
 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 
haftmann 
parents: 
59083 
diff
changeset
 | 
236  | 
let  | 
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
237  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
238  | 
fun matches obj ctxt' = Pattern.matches thy (pat, obj) orelse  | 
| 
59095
 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 
haftmann 
parents: 
59083 
diff
changeset
 | 
239  | 
(case obj of  | 
| 77995 | 240  | 
Abs _ =>  | 
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
241  | 
let val ((_, t'), ctxt'') = Variable.dest_abs obj ctxt'  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
242  | 
in matches t' ctxt'' end  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
243  | 
| t $ u => matches t ctxt' orelse matches u ctxt'  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
244  | 
| _ => false);  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
245  | 
in matches obj ctxt end;  | 
| 
59095
 
3100a7b1c092
turn application-specific Pattern.matches_subterm into an application-private function
 
haftmann 
parents: 
59083 
diff
changeset
 | 
246  | 
|
| 52940 | 247  | 
(*Including all constants and frees is only sound because matching  | 
248  | 
uses higher-order patterns. If full matching were used, then  | 
|
249  | 
constants that may be subject to beta-reduction after substitution  | 
|
250  | 
of frees should not be included for LHS set because they could be  | 
|
251  | 
thrown away by the substituted function. E.g. for (?F 1 2) do not  | 
|
| 67721 | 252  | 
include 1 or 2, if it were possible for ?F to be (\<lambda>x y. 3). The  | 
| 52940 | 253  | 
largest possible set should always be included on the RHS.*)  | 
| 30143 | 254  | 
|
255  | 
fun filter_pattern ctxt pat =  | 
|
256  | 
let  | 
|
| 
59098
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
257  | 
val pat' = (expand_abs o Envir.eta_contract) pat;  | 
| 
 
b6ba3adb48e3
eta-expand all search patterns using schematic place holders
 
haftmann 
parents: 
59097 
diff
changeset
 | 
258  | 
val pat_consts = get_names pat';  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
259  | 
fun check ((x, thm), NONE) = check ((x, thm), SOME (get_names (Thm.full_prop_of thm)))  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
260  | 
| check ((_, thm), c as SOME thm_consts) =  | 
| 33038 | 261  | 
(if subset (op =) (pat_consts, thm_consts) andalso  | 
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
70545 
diff
changeset
 | 
262  | 
matches_subterm ctxt (pat', Thm.full_prop_of thm)  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
263  | 
then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0) else NONE, c);  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
264  | 
in check end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
265  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
266  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
267  | 
(* interpret criteria as filters *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
|
| 16036 | 269  | 
local  | 
270  | 
||
271  | 
fun err_no_goal c =  | 
|
272  | 
  error ("Current goal required for " ^ c ^ " search criterion");
 | 
|
273  | 
||
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
274  | 
fun filter_crit _ _ (Name name) = apfst (filter_name name)  | 
| 16036 | 275  | 
| filter_crit _ NONE Intro = err_no_goal "intro"  | 
276  | 
| filter_crit _ NONE Elim = err_no_goal "elim"  | 
|
277  | 
| filter_crit _ NONE Dest = err_no_goal "dest"  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
278  | 
| filter_crit _ NONE Solves = err_no_goal "solves"  | 
| 52940 | 279  | 
| filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt (Thm.prop_of goal))  | 
280  | 
| filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (Thm.prop_of goal))  | 
|
281  | 
| filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (Thm.prop_of goal))  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
282  | 
| filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal)  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
283  | 
| filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat)  | 
| 16088 | 284  | 
| filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;  | 
| 16036 | 285  | 
|
| 53632 | 286  | 
fun opt_not x = if is_some x then NONE else SOME (0, 0, 0);  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
287  | 
|
| 53632 | 288  | 
fun opt_add (SOME (a, c, x)) (SOME (b, d, y)) = SOME (Int.max (a,b), Int.max (c, d), x + y : int)  | 
| 26283 | 289  | 
| opt_add _ _ = NONE;  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
290  | 
|
| 30143 | 291  | 
fun app_filters thm =  | 
292  | 
let  | 
|
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
293  | 
fun app (NONE, _, _) = NONE  | 
| 32798 | 294  | 
| app (SOME v, _, []) = SOME (v, thm)  | 
| 30143 | 295  | 
| app (r, consts, f :: fs) =  | 
296  | 
let val (r', consts') = f (thm, consts)  | 
|
297  | 
in app (opt_add r r', consts', fs) end;  | 
|
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
298  | 
in app end;  | 
| 
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
299  | 
|
| 16036 | 300  | 
in  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
301  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
302  | 
fun filter_criterion ctxt opt_goal (b, c) =  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
303  | 
(if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c;  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
304  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
305  | 
fun sorted_filter filters thms =  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
306  | 
let  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
307  | 
fun eval_filters thm = app_filters thm (SOME (0, 0, 0), NONE, filters);  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
308  | 
|
| 53632 | 309  | 
(*filters return: (thm size, number of assumptions, substitution size) option, so  | 
310  | 
sort according to size of thm first, then number of assumptions,  | 
|
311  | 
then by the substitution size, then by term order *)  | 
|
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
312  | 
fun result_ord (((p0, s0, t0), (_, thm0)), ((p1, s1, t1), (_, thm1))) =  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
313  | 
prod_ord int_ord (prod_ord int_ord (prod_ord int_ord Term_Ord.term_ord))  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
314  | 
((p1, (s1, (t1, Thm.full_prop_of thm1))), (p0, (s0, (t0, Thm.full_prop_of thm0))));  | 
| 46977 | 315  | 
in  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
316  | 
grouped 100 Par_List.map eval_filters thms  | 
| 46977 | 317  | 
|> map_filter I |> sort result_ord |> map #2  | 
318  | 
end;  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
319  | 
|
| 30822 | 320  | 
fun lazy_filter filters =  | 
321  | 
let  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
322  | 
fun lazy_match thms = Seq.make (fn () => first_match thms)  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
323  | 
and first_match [] = NONE  | 
| 30822 | 324  | 
| first_match (thm :: thms) =  | 
| 53632 | 325  | 
(case app_filters thm (SOME (0, 0, 0), NONE, filters) of  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
326  | 
NONE => first_match thms  | 
| 30822 | 327  | 
| SOME (_, t) => SOME (t, lazy_match thms));  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
328  | 
in lazy_match end;  | 
| 30822 | 329  | 
|
| 16036 | 330  | 
end;  | 
331  | 
||
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
|
| 52940 | 333  | 
(* removing duplicates, preferring nicer names, roughly O(n log n) *)  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
334  | 
|
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
335  | 
local  | 
| 
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
336  | 
|
| 59916 | 337  | 
val hidden_ord = bool_ord o apply2 Long_Name.is_hidden;  | 
| 77854 | 338  | 
val qual_ord = int_ord o apply2 Long_Name.count;  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
339  | 
val txt_ord = int_ord o apply2 size;  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
340  | 
|
| 
63080
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
341  | 
fun nicer_name ((a, x), i) ((b, y), j) =  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
342  | 
(case bool_ord (a, b) of EQUAL =>  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
343  | 
(case hidden_ord (x, y) of EQUAL =>  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
344  | 
(case int_ord (i, j) of EQUAL =>  | 
| 
63080
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
345  | 
(case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord)  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
346  | 
| ord => ord)  | 
| 27486 | 347  | 
| ord => ord)  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
348  | 
| ord => ord) <> GREATER;  | 
| 
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
349  | 
|
| 29848 | 350  | 
fun rem_cdups nicer xs =  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
351  | 
let  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
352  | 
fun rem_c rev_seen [] = rev rev_seen  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
353  | 
| rem_c rev_seen [x] = rem_c (x :: rev_seen) []  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
354  | 
| rem_c rev_seen ((x as ((n, thm), _)) :: (y as ((n', thm'), _)) :: rest) =  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
355  | 
if Thm.eq_thm_prop (thm, thm')  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
356  | 
then rem_c rev_seen ((if nicer n n' then x else y) :: rest)  | 
| 
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
357  | 
else rem_c (x :: rev_seen) (y :: rest);  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
358  | 
in rem_c [] xs end;  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
359  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
360  | 
in  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
361  | 
|
| 30143 | 362  | 
fun nicer_shortest ctxt =  | 
363  | 
let  | 
|
| 
56143
 
ed2b660a52a1
more accurate resolution of hybrid facts, which actually changes the sort order of results;
 
wenzelm 
parents: 
56141 
diff
changeset
 | 
364  | 
fun extern_shortest name =  | 
| 
63080
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
365  | 
let  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
366  | 
val facts = Proof_Context.facts_of_fact ctxt name;  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
367  | 
val space = Facts.space_of facts;  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
368  | 
in (Facts.is_dynamic facts name, Name_Space.extern_shortest ctxt space name) end;  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
369  | 
in fn (x, i) => fn (y, j) => nicer_name (extern_shortest x, i) (extern_shortest y, j) end;  | 
| 29848 | 370  | 
|
371  | 
fun rem_thm_dups nicer xs =  | 
|
| 52940 | 372  | 
(xs ~~ (1 upto length xs))  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
373  | 
|> sort (Term_Ord.fast_term_ord o apply2 (Thm.full_prop_of o #2 o #1))  | 
| 29848 | 374  | 
|> rem_cdups nicer  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
375  | 
|> sort (int_ord o apply2 #2)  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
376  | 
|> map #1;  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
377  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
378  | 
end;  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
379  | 
|
| 
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
380  | 
|
| 52941 | 381  | 
|
382  | 
(** main operations **)  | 
|
383  | 
||
384  | 
(* filter_theorems *)  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
385  | 
|
| 26283 | 386  | 
fun all_facts_of ctxt =  | 
| 33381 | 387  | 
let  | 
| 61054 | 388  | 
val thy = Proof_Context.theory_of ctxt;  | 
389  | 
val transfer = Global_Theory.transfer_theories thy;  | 
|
| 
56158
 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 
wenzelm 
parents: 
56143 
diff
changeset
 | 
390  | 
val local_facts = Proof_Context.facts_of ctxt;  | 
| 61054 | 391  | 
val global_facts = Global_Theory.facts_of thy;  | 
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
392  | 
in  | 
| 
63080
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
393  | 
(Facts.dest_all (Context.Proof ctxt) false [global_facts] local_facts @  | 
| 
 
8326aa594273
find dynamic facts as well, but static ones are preferred;
 
wenzelm 
parents: 
62969 
diff
changeset
 | 
394  | 
Facts.dest_all (Context.Proof ctxt) false [] global_facts)  | 
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
395  | 
|> maps Thm_Name.make_list  | 
| 61054 | 396  | 
|> map (apsnd transfer)  | 
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
397  | 
end;  | 
| 17972 | 398  | 
|
| 43070 | 399  | 
fun filter_theorems ctxt theorems query =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
400  | 
let  | 
| 46718 | 401  | 
    val {goal = opt_goal, limit = opt_limit, rem_dups, criteria} = query;
 | 
| 
43069
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
402  | 
val filters = map (filter_criterion ctxt opt_goal) criteria;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
403  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
404  | 
fun find_all theorems =  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
405  | 
let  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
406  | 
val raw_matches = sorted_filter filters theorems;  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
407  | 
|
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
408  | 
val matches =  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
409  | 
if rem_dups  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
410  | 
then rem_thm_dups (nicer_shortest ctxt) raw_matches  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
411  | 
else raw_matches;  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
412  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
413  | 
val len = length matches;  | 
| 67147 | 414  | 
val lim = the_default (Options.default_int \<^system_option>\<open>find_theorems_limit\<close>) opt_limit;  | 
| 34088 | 415  | 
in (SOME len, drop (Int.max (len - lim, 0)) matches) end;  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
416  | 
|
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
417  | 
val find =  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
418  | 
if rem_dups orelse is_none opt_limit  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
419  | 
then find_all  | 
| 30822 | 420  | 
else pair NONE o Seq.list_of o Seq.take (the opt_limit) o lazy_filter filters;  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
421  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
422  | 
in find theorems end;  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
423  | 
|
| 46718 | 424  | 
fun filter_theorems_cmd ctxt theorems raw_query =  | 
| 52941 | 425  | 
filter_theorems ctxt theorems (map_criteria (map (apsnd (read_criterion ctxt))) raw_query);  | 
426  | 
||
427  | 
||
428  | 
(* find_theorems *)  | 
|
429  | 
||
430  | 
local  | 
|
| 43067 | 431  | 
|
432  | 
fun gen_find_theorems filter ctxt opt_goal opt_limit rem_dups raw_criteria =  | 
|
| 
43069
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
433  | 
let  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
434  | 
val assms =  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
435  | 
Proof_Context.get_fact ctxt (Facts.named "local.assms")  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
436  | 
handle ERROR _ => [];  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
437  | 
val add_prems = Seq.hd o TRY (Method.insert_tac ctxt assms 1);  | 
| 
43069
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
438  | 
val opt_goal' = Option.map add_prems opt_goal;  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
439  | 
in  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
440  | 
filter ctxt (all_facts_of ctxt)  | 
| 46718 | 441  | 
      {goal = opt_goal', limit = opt_limit, rem_dups = rem_dups, criteria = raw_criteria}
 | 
| 
43069
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
442  | 
end;  | 
| 
30186
 
1f836e949ac2
replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
 
wenzelm 
parents: 
30143 
diff
changeset
 | 
443  | 
|
| 52941 | 444  | 
in  | 
445  | 
||
| 43067 | 446  | 
val find_theorems = gen_find_theorems filter_theorems;  | 
447  | 
val find_theorems_cmd = gen_find_theorems filter_theorems_cmd;  | 
|
448  | 
||
| 52941 | 449  | 
end;  | 
450  | 
||
451  | 
||
452  | 
(* pretty_theorems *)  | 
|
453  | 
||
454  | 
local  | 
|
455  | 
||
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
456  | 
fun pretty_thm_head ctxt (name, i) =  | 
| 
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
457  | 
[Pretty.marks_str (#1 (Proof_Context.markup_extern_fact ctxt name), name),  | 
| 
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
458  | 
Pretty.str (Thm_Name.print_suffix (name, i)), Pretty.str ":", Pretty.brk 1];  | 
| 49888 | 459  | 
|
| 52941 | 460  | 
in  | 
461  | 
||
| 
80330
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
462  | 
fun pretty_thm ctxt (thm_name, thm) =  | 
| 
 
e01aae620437
clarified signature: prefer internal Thm_Name.T over external Facts.ref;
 
wenzelm 
parents: 
80328 
diff
changeset
 | 
463  | 
Pretty.block (pretty_thm_head ctxt thm_name @ [Thm.pretty_thm ctxt thm]);  | 
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
464  | 
|
| 52941 | 465  | 
fun pretty_theorems state opt_limit rem_dups raw_criteria =  | 
| 30143 | 466  | 
let  | 
| 52941 | 467  | 
val ctxt = Proof.context_of state;  | 
| 64984 | 468  | 
val opt_goal = try (#goal o Proof.simple_goal) state;  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
469  | 
val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;  | 
| 52941 | 470  | 
|
| 52940 | 471  | 
val (opt_found, theorems) =  | 
| 
55671
 
aeca05e62fef
removed remains of old experiment (see b933142e02d0);
 
wenzelm 
parents: 
55670 
diff
changeset
 | 
472  | 
filter_theorems ctxt (all_facts_of ctxt)  | 
| 52855 | 473  | 
        {goal = opt_goal, limit = opt_limit, rem_dups = rem_dups, criteria = criteria};
 | 
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
474  | 
val returned = length theorems;  | 
| 
31684
 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 
wenzelm 
parents: 
31042 
diff
changeset
 | 
475  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
476  | 
val tally_msg =  | 
| 52940 | 477  | 
(case opt_found of  | 
| 38335 | 478  | 
NONE => "displaying " ^ string_of_int returned ^ " theorem(s)"  | 
| 30822 | 479  | 
| SOME found =>  | 
| 38335 | 480  | 
"found " ^ string_of_int found ^ " theorem(s)" ^  | 
| 30822 | 481  | 
(if returned < found  | 
482  | 
             then " (" ^ string_of_int returned ^ " displayed)"
 | 
|
483  | 
else ""));  | 
|
| 
56912
 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 
wenzelm 
parents: 
56908 
diff
changeset
 | 
484  | 
val position_markup = Position.markup (Position.thread_data ()) Markup.position;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
485  | 
in  | 
| 
56891
 
48899c43b07d
tuned message -- more context for detached window etc.;
 
wenzelm 
parents: 
56621 
diff
changeset
 | 
486  | 
Pretty.block  | 
| 
56912
 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 
wenzelm 
parents: 
56908 
diff
changeset
 | 
487  | 
(Pretty.fbreaks  | 
| 
 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 
wenzelm 
parents: 
56908 
diff
changeset
 | 
488  | 
(Pretty.mark position_markup (Pretty.keyword1 "find_theorems") ::  | 
| 
 
293cd4dcfebc
some position markup to help locating the query context, e.g. from "Info" dockable;
 
wenzelm 
parents: 
56908 
diff
changeset
 | 
489  | 
map (pretty_criterion ctxt) criteria)) ::  | 
| 38335 | 490  | 
Pretty.str "" ::  | 
| 
56908
 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 
wenzelm 
parents: 
56891 
diff
changeset
 | 
491  | 
(if null theorems then [Pretty.str "found nothing"]  | 
| 38335 | 492  | 
else  | 
| 
56908
 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 
wenzelm 
parents: 
56891 
diff
changeset
 | 
493  | 
Pretty.str (tally_msg ^ ":") ::  | 
| 
 
734f7e6151c9
tuned message: more compact, imitate actual command line;
 
wenzelm 
parents: 
56891 
diff
changeset
 | 
494  | 
grouped 10 Par_List.map (Pretty.item o single o pretty_thm ctxt) (rev theorems))  | 
| 80328 | 495  | 
end |> Pretty.fbreaks |> Pretty.block0;  | 
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
496  | 
|
| 52941 | 497  | 
end;  | 
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
498  | 
|
| 32798 | 499  | 
|
| 46718 | 500  | 
|
| 
52865
 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 
wenzelm 
parents: 
52863 
diff
changeset
 | 
501  | 
(** Isar command syntax **)  | 
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
502  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
503  | 
local  | 
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
504  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
505  | 
val criterion =  | 
| 62969 | 506  | 
Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.name) >> Name ||  | 
| 36950 | 507  | 
Parse.reserved "intro" >> K Intro ||  | 
508  | 
Parse.reserved "elim" >> K Elim ||  | 
|
509  | 
Parse.reserved "dest" >> K Dest ||  | 
|
510  | 
Parse.reserved "solves" >> K Solves ||  | 
|
511  | 
Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp ||  | 
|
512  | 
Parse.term >> Pattern;  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
513  | 
|
| 74555 | 514  | 
val query_keywords = Keyword.add_minor_keywords [":"] Keyword.empty_keywords;  | 
| 58905 | 515  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
516  | 
in  | 
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
517  | 
|
| 
62848
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
518  | 
val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion);  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
519  | 
|
| 52925 | 520  | 
fun read_query pos str =  | 
| 59083 | 521  | 
Token.explode query_keywords pos str  | 
| 52855 | 522  | 
|> filter Token.is_proper  | 
| 
62848
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
523  | 
|> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof)))  | 
| 52925 | 524  | 
|> #1;  | 
| 
43068
 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 
krauss 
parents: 
43067 
diff
changeset
 | 
525  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
526  | 
end;  | 
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
527  | 
|
| 
52851
 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 
wenzelm 
parents: 
52788 
diff
changeset
 | 
528  | 
|
| 
 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 
wenzelm 
parents: 
52788 
diff
changeset
 | 
529  | 
|
| 
52865
 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 
wenzelm 
parents: 
52863 
diff
changeset
 | 
530  | 
(** PIDE query operation **)  | 
| 
52854
 
92932931bd82
more general Output.result: allow to update arbitrary properties;
 
wenzelm 
parents: 
52851 
diff
changeset
 | 
531  | 
|
| 
62848
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
532  | 
fun proof_state st =  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
533  | 
(case try Toplevel.proof_of st of  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
534  | 
SOME state => state  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
535  | 
| NONE => Proof.init (Toplevel.context_of st));  | 
| 
 
e4140efe699e
clarified bootstrap -- more uniform use of ML files;
 
wenzelm 
parents: 
61841 
diff
changeset
 | 
536  | 
|
| 
52865
 
02a7e7180ee5
slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
 
wenzelm 
parents: 
52863 
diff
changeset
 | 
537  | 
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: 
59970 
diff
changeset
 | 
538  | 
  Query_Operation.register {name = "find_theorems", pri = Task_Queue.urgent_pri}
 | 
| 61223 | 539  | 
    (fn {state = st, 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: 
59970 
diff
changeset
 | 
540  | 
if can Toplevel.context_of st then  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
541  | 
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: 
59970 
diff
changeset
 | 
542  | 
val [limit_arg, allow_dups_arg, 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: 
59970 
diff
changeset
 | 
543  | 
val state = proof_state st;  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
544  | 
val opt_limit = Int.fromString limit_arg;  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
545  | 
val rem_dups = allow_dups_arg = "false";  | 
| 
 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
546  | 
val criteria = read_query Position.none query_arg;  | 
| 61223 | 547  | 
in writeln_result (Pretty.string_of (pretty_theorems state opt_limit rem_dups criteria)) 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: 
59970 
diff
changeset
 | 
548  | 
else error "Unknown context");  | 
| 
52851
 
e71b5160f242
minimal print function "find_theorems", which merely echos its arguments;
 
wenzelm 
parents: 
52788 
diff
changeset
 | 
549  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
550  | 
end;  |