| author | wenzelm | 
| Tue, 03 Apr 2012 16:51:01 +0200 | |
| changeset 47294 | 008b7858f3c0 | 
| parent 46977 | bd0ee92cabe7 | 
| child 48646 | 91281e9472d8 | 
| 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  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
12  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
13  | 
datatype theorem =  | 
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
14  | 
Internal of Facts.ref * thm | External of Facts.ref * term  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
15  | 
|
| 43070 | 16  | 
  type 'term query = {
 | 
17  | 
goal: thm option,  | 
|
18  | 
limit: int option,  | 
|
19  | 
rem_dups: bool,  | 
|
20  | 
criteria: (bool * 'term criterion) list  | 
|
21  | 
}  | 
|
22  | 
||
| 32738 | 23  | 
val tac_limit: int Unsynchronized.ref  | 
24  | 
val limit: int Unsynchronized.ref  | 
|
| 43067 | 25  | 
|
26  | 
val read_criterion: Proof.context -> string criterion -> term criterion  | 
|
| 
43068
 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 
krauss 
parents: 
43067 
diff
changeset
 | 
27  | 
val query_parser: (bool * string criterion) list parser  | 
| 43067 | 28  | 
|
| 43071 | 29  | 
val xml_of_query: term query -> XML.tree  | 
30  | 
val query_of_xml: XML.tree -> term query  | 
|
31  | 
val xml_of_result: int option * theorem list -> XML.tree  | 
|
32  | 
val result_of_xml: XML.tree -> int option * theorem list  | 
|
33  | 
||
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
34  | 
val find_theorems: Proof.context -> thm option -> int option -> bool ->  | 
| 43067 | 35  | 
(bool * term criterion) list -> int option * (Facts.ref * thm) list  | 
36  | 
val find_theorems_cmd: Proof.context -> thm option -> int option -> bool ->  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
37  | 
(bool * string criterion) list -> int option * (Facts.ref * thm) list  | 
| 43070 | 38  | 
|
39  | 
val filter_theorems: Proof.context -> theorem list -> term query ->  | 
|
| 43067 | 40  | 
int option * theorem list  | 
| 43070 | 41  | 
val filter_theorems_cmd: Proof.context -> theorem list -> string query ->  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
42  | 
int option * theorem list  | 
| 
41841
 
c27b0b37041a
Refactor find_theorems to provide a more general filter_facts method
 
noschinl 
parents: 
41835 
diff
changeset
 | 
43  | 
|
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
44  | 
val pretty_theorem: Proof.context -> theorem -> Pretty.T  | 
| 
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
 | 
45  | 
val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T  | 
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
46  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
|
| 33301 | 49  | 
structure Find_Theorems: FIND_THEOREMS =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
struct  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
(** search criteria **)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
|
| 16036 | 54  | 
datatype 'term criterion =  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
55  | 
Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term;  | 
| 16036 | 56  | 
|
| 
33036
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
57  | 
fun apply_dummies tm =  | 
| 33301 | 58  | 
let  | 
59  | 
val (xs, _) = Term.strip_abs tm;  | 
|
60  | 
val tm' = Term.betapplys (tm, map (Term.dummy_pattern o #2) xs);  | 
|
61  | 
in #1 (Term.replace_dummy_patterns tm' 1) end;  | 
|
| 
33036
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
62  | 
|
| 
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
63  | 
fun parse_pattern ctxt nm =  | 
| 
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
64  | 
let  | 
| 42360 | 65  | 
val consts = Proof_Context.consts_of ctxt;  | 
| 33301 | 66  | 
val nm' =  | 
67  | 
(case Syntax.parse_term ctxt nm of  | 
|
68  | 
Const (c, _) => c  | 
|
69  | 
| _ => Consts.intern consts nm);  | 
|
| 
33036
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
70  | 
in  | 
| 33301 | 71  | 
(case try (Consts.the_abbreviation consts) nm' of  | 
| 42360 | 72  | 
SOME (_, rhs) => apply_dummies (Proof_Context.expand_abbrevs ctxt rhs)  | 
73  | 
| NONE => Proof_Context.read_term_pattern ctxt nm)  | 
|
| 
33036
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
74  | 
end;  | 
| 
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
75  | 
|
| 16036 | 76  | 
fun read_criterion _ (Name name) = Name name  | 
77  | 
| read_criterion _ Intro = Intro  | 
|
78  | 
| read_criterion _ Elim = Elim  | 
|
79  | 
| read_criterion _ Dest = Dest  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
80  | 
| read_criterion _ Solves = Solves  | 
| 42360 | 81  | 
| read_criterion ctxt (Simp str) = Simp (Proof_Context.read_term_pattern ctxt str)  | 
| 
33036
 
c61fe520602b
find_theorems: better handling of abbreviations (by Timothy Bourke)
 
kleing 
parents: 
33029 
diff
changeset
 | 
82  | 
| read_criterion ctxt (Pattern str) = Pattern (parse_pattern ctxt str);  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
|
| 16036 | 84  | 
fun pretty_criterion ctxt (b, c) =  | 
85  | 
let  | 
|
86  | 
fun prfx s = if b then s else "-" ^ s;  | 
|
87  | 
in  | 
|
88  | 
(case c of  | 
|
89  | 
Name name => Pretty.str (prfx "name: " ^ quote name)  | 
|
90  | 
| Intro => Pretty.str (prfx "intro")  | 
|
91  | 
| Elim => Pretty.str (prfx "elim")  | 
|
92  | 
| 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
 | 
93  | 
| Solves => Pretty.str (prfx "solves")  | 
| 16088 | 94  | 
| Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,  | 
| 24920 | 95  | 
Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]  | 
| 16036 | 96  | 
| Pattern pat => Pretty.enclose (prfx " \"") "\""  | 
| 24920 | 97  | 
[Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])  | 
| 16036 | 98  | 
end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
|
| 
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
 | 
100  | 
|
| 43620 | 101  | 
|
| 43070 | 102  | 
(** queries **)  | 
103  | 
||
104  | 
type 'term query = {
 | 
|
105  | 
goal: thm option,  | 
|
106  | 
limit: int option,  | 
|
107  | 
rem_dups: bool,  | 
|
108  | 
criteria: (bool * 'term criterion) list  | 
|
109  | 
};  | 
|
110  | 
||
111  | 
fun map_criteria f {goal, limit, rem_dups, criteria} =
 | 
|
| 46718 | 112  | 
  {goal = goal, limit = limit, rem_dups = rem_dups, criteria = f criteria};
 | 
| 43070 | 113  | 
|
| 43071 | 114  | 
fun xml_of_criterion (Name name) = XML.Elem (("Name", [("val", name)]), [])
 | 
115  | 
  | xml_of_criterion Intro = XML.Elem (("Intro", []) , [])
 | 
|
116  | 
  | xml_of_criterion Elim = XML.Elem (("Elim", []) , [])
 | 
|
117  | 
  | xml_of_criterion Dest = XML.Elem (("Dest", []) , [])
 | 
|
118  | 
  | xml_of_criterion Solves = XML.Elem (("Solves", []) , [])
 | 
|
119  | 
  | xml_of_criterion (Simp pat) = XML.Elem (("Simp", []) , [XML_Syntax.xml_of_term pat])
 | 
|
120  | 
  | xml_of_criterion (Pattern pat) = XML.Elem (("Pattern", []) , [XML_Syntax.xml_of_term pat]);
 | 
|
121  | 
||
122  | 
fun criterion_of_xml (XML.Elem (("Name", [("val", name)]), [])) = Name name
 | 
|
123  | 
  | criterion_of_xml (XML.Elem (("Intro", []) , [])) = Intro
 | 
|
124  | 
  | criterion_of_xml (XML.Elem (("Elim", []) , [])) = Elim
 | 
|
125  | 
  | criterion_of_xml (XML.Elem (("Dest", []) , [])) = Dest
 | 
|
126  | 
  | criterion_of_xml (XML.Elem (("Solves", []) , [])) = Solves
 | 
|
127  | 
  | criterion_of_xml (XML.Elem (("Simp", []) , [tree])) = Simp (XML_Syntax.term_of_xml tree)
 | 
|
128  | 
  | criterion_of_xml (XML.Elem (("Pattern", []) , [tree])) = Pattern (XML_Syntax.term_of_xml tree)
 | 
|
129  | 
  | criterion_of_xml tree = raise XML_Syntax.XML ("criterion_of_xml: bad tree", tree);
 | 
|
130  | 
||
| 46718 | 131  | 
fun xml_of_query {goal = NONE, limit, rem_dups, criteria} =
 | 
| 43071 | 132  | 
let  | 
133  | 
val properties = []  | 
|
134  | 
          |> (if rem_dups then cons ("rem_dups", "") else I)
 | 
|
135  | 
          |> (if is_some limit then cons ("limit", Markup.print_int (the limit)) else I);
 | 
|
136  | 
in  | 
|
| 43767 | 137  | 
        XML.Elem (("Query", properties), XML.Encode.list
 | 
138  | 
(XML.Encode.pair XML.Encode.bool (single o xml_of_criterion)) criteria)  | 
|
| 43071 | 139  | 
end  | 
140  | 
| xml_of_query _ = raise Fail "cannot serialize goal";  | 
|
141  | 
||
142  | 
fun query_of_xml (XML.Elem (("Query", properties), body)) =
 | 
|
143  | 
let  | 
|
144  | 
val rem_dups = Properties.defined properties "rem_dups";  | 
|
145  | 
val limit = Properties.get properties "limit" |> Option.map Markup.parse_int;  | 
|
| 43724 | 146  | 
val criteria =  | 
| 43767 | 147  | 
XML.Decode.list (XML.Decode.pair XML.Decode.bool  | 
| 43724 | 148  | 
(criterion_of_xml o the_single)) body;  | 
| 43071 | 149  | 
in  | 
| 46718 | 150  | 
        {goal = NONE, limit = limit, rem_dups = rem_dups, criteria = criteria}
 | 
| 43071 | 151  | 
end  | 
152  | 
  | query_of_xml tree = raise XML_Syntax.XML ("query_of_xml: bad tree", tree);
 | 
|
| 43070 | 153  | 
|
| 43620 | 154  | 
|
155  | 
||
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
156  | 
(** theorems, either internal or external (without proof) **)  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
157  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
158  | 
datatype theorem =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
159  | 
Internal of Facts.ref * thm |  | 
| 43071 | 160  | 
External of Facts.ref * term; (* FIXME: Facts.ref not appropriate *)  | 
161  | 
||
162  | 
fun fact_ref_markup (Facts.Named ((name, pos), SOME [Facts.Single i])) =  | 
|
163  | 
      Position.markup pos o Markup.properties [("name", name), ("index", Markup.print_int i)]
 | 
|
164  | 
| fact_ref_markup (Facts.Named ((name, pos), NONE)) =  | 
|
165  | 
      Position.markup pos o Markup.properties [("name", name)]
 | 
|
| 43620 | 166  | 
| fact_ref_markup fact_ref = raise Fail "bad fact ref";  | 
| 43071 | 167  | 
|
168  | 
fun xml_of_theorem (Internal _) = raise Fail "xml_of_theorem: Internal"  | 
|
169  | 
| xml_of_theorem (External (fact_ref, prop)) =  | 
|
| 43620 | 170  | 
      XML.Elem (fact_ref_markup fact_ref ("External", []), [XML_Syntax.xml_of_term prop]);
 | 
| 43071 | 171  | 
|
172  | 
fun theorem_of_xml (XML.Elem (("External", properties), [tree])) =
 | 
|
| 43620 | 173  | 
let  | 
174  | 
val name = the (Properties.get properties "name");  | 
|
175  | 
val pos = Position.of_properties properties;  | 
|
| 46718 | 176  | 
val intvs_opt =  | 
177  | 
Option.map (single o Facts.Single o Markup.parse_int)  | 
|
178  | 
(Properties.get properties "index");  | 
|
| 43620 | 179  | 
in  | 
180  | 
External (Facts.Named ((name, pos), intvs_opt), XML_Syntax.term_of_xml tree)  | 
|
181  | 
end  | 
|
| 43071 | 182  | 
  | theorem_of_xml tree = raise XML_Syntax.XML ("theorem_of_xml: bad tree", tree);
 | 
183  | 
||
184  | 
fun xml_of_result (opt_found, theorems) =  | 
|
185  | 
let  | 
|
186  | 
val properties =  | 
|
187  | 
      if is_some opt_found then [("found", Markup.print_int (the opt_found))] else [];
 | 
|
188  | 
in  | 
|
| 43767 | 189  | 
    XML.Elem (("Result", properties), XML.Encode.list (single o xml_of_theorem) theorems)
 | 
| 43071 | 190  | 
end;  | 
191  | 
||
192  | 
fun result_of_xml (XML.Elem (("Result", properties), body)) =
 | 
|
193  | 
(Properties.get properties "found" |> Option.map Markup.parse_int,  | 
|
| 43767 | 194  | 
XML.Decode.list (theorem_of_xml o the_single) body)  | 
| 43071 | 195  | 
  | result_of_xml tree = raise XML_Syntax.XML ("result_of_xml: bad tree", tree);
 | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
196  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
197  | 
fun prop_of (Internal (_, thm)) = Thm.full_prop_of thm  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
198  | 
| prop_of (External (_, prop)) = prop;  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
199  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
200  | 
fun nprems_of (Internal (_, thm)) = Thm.nprems_of thm  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
201  | 
| nprems_of (External (_, prop)) = Logic.count_prems prop;  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
202  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
203  | 
fun major_prem_of (Internal (_, thm)) = Thm.major_prem_of thm  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
204  | 
| major_prem_of (External (_, prop)) =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
205  | 
Logic.strip_assums_concl (hd (Logic.strip_imp_prems prop));  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
206  | 
|
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
207  | 
fun fact_ref_of (Internal (fact_ref, _)) = fact_ref  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
208  | 
| fact_ref_of (External (fact_ref, _)) = fact_ref;  | 
| 
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
 | 
209  | 
|
| 43620 | 210  | 
|
211  | 
||
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
(** search criterion filters **)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
214  | 
(*generated filters are to be of the form  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
215  | 
input: theorem  | 
| 17106 | 216  | 
output: (p:int, s:int) option, where  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
217  | 
NONE indicates no match  | 
| 17106 | 218  | 
p is the primary sorting criterion  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
219  | 
(eg. number of assumptions in the theorem)  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
220  | 
s is the secondary sorting criterion  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
221  | 
(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
 | 
222  | 
when applying a set of filters to a thm, fold results in:  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
223  | 
(biggest p, sum of all s)  | 
| 17106 | 224  | 
currently p and s only matter for intro, elim, dest and simp filters,  | 
225  | 
otherwise the default ordering is used.  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
226  | 
*)  | 
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
227  | 
|
| 16088 | 228  | 
|
229  | 
(* matching theorems *)  | 
|
| 17106 | 230  | 
|
| 35625 | 231  | 
fun is_nontrivial thy = Term.is_Const o Term.head_of o Object_Logic.drop_judgment thy;  | 
| 16088 | 232  | 
|
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
233  | 
(*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
 | 
234  | 
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
 | 
235  | 
trivial matches are ignored.  | 
| 
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
236  | 
returns: smallest substitution size*)  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
237  | 
fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =  | 
| 16088 | 238  | 
let  | 
| 42360 | 239  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 16088 | 240  | 
|
| 16486 | 241  | 
fun matches pat =  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
242  | 
is_nontrivial thy pat andalso  | 
| 
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
243  | 
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
 | 
244  | 
|
| 
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
245  | 
fun substsize pat =  | 
| 18184 | 246  | 
let val (_, subst) =  | 
247  | 
Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)  | 
|
| 17205 | 248  | 
in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;  | 
| 16088 | 249  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
250  | 
fun bestmatch [] = NONE  | 
| 33029 | 251  | 
| bestmatch xs = SOME (foldl1 Int.min xs);  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
252  | 
|
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
253  | 
val match_thm = matches o refine_term;  | 
| 16486 | 254  | 
in  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
255  | 
map (substsize o refine_term) (filter match_thm (extract_terms term_src))  | 
| 26283 | 256  | 
|> bestmatch  | 
| 16088 | 257  | 
end;  | 
258  | 
||
259  | 
||
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
260  | 
(* filter_name *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
262  | 
fun filter_name str_pat theorem =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
263  | 
if match_string str_pat (Facts.name_of_ref (fact_ref_of theorem))  | 
| 17205 | 264  | 
then SOME (0, 0) else NONE;  | 
| 
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  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
267  | 
(* filter intro/elim/dest/solves rules *)  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
269  | 
fun filter_dest ctxt goal theorem =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
let  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
271  | 
val extract_dest =  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
272  | 
(fn theorem => if nprems_of theorem = 0 then [] else [prop_of theorem],  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
273  | 
hd o Logic.strip_imp_prems);  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
274  | 
val prems = Logic.prems_of_goal goal 1;  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
275  | 
|
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
276  | 
fun try_subst prem = is_matching_thm extract_dest ctxt true prem theorem;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19476 
diff
changeset
 | 
277  | 
val successful = prems |> map_filter try_subst;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
278  | 
in  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
279  | 
(*if possible, keep best substitution (one with smallest size)*)  | 
| 17106 | 280  | 
(*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
 | 
281  | 
assumption is as good as an intro rule with none*)  | 
| 17205 | 282  | 
if not (null successful)  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
283  | 
then SOME (nprems_of theorem - 1, foldl1 Int.min successful) else NONE  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
284  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
285  | 
|
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
286  | 
fun filter_intro ctxt goal theorem =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
287  | 
let  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
288  | 
val extract_intro = (single o prop_of, Logic.strip_imp_concl);  | 
| 16036 | 289  | 
val concl = Logic.concl_of_goal goal 1;  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
290  | 
val ss = is_matching_thm extract_intro ctxt true concl theorem;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
291  | 
in  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
292  | 
if is_some ss then SOME (nprems_of theorem, the ss) else NONE  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
293  | 
end;  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
294  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
295  | 
fun filter_elim ctxt goal theorem =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
296  | 
if nprems_of theorem > 0 then  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
297  | 
let  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
298  | 
val rule = prop_of theorem;  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
299  | 
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
 | 
300  | 
val goal_concl = Logic.concl_of_goal goal 1;  | 
| 26283 | 301  | 
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
 | 
302  | 
val rule_concl = Logic.strip_imp_concl rule;  | 
| 26283 | 303  | 
      fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);
 | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
304  | 
val rule_tree = combine rule_mp rule_concl;  | 
| 26283 | 305  | 
fun goal_tree prem = combine prem goal_concl;  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
306  | 
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
 | 
307  | 
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
 | 
308  | 
in  | 
| 32798 | 309  | 
(*elim rules always have assumptions, so an elim with one  | 
310  | 
assumption is as good as an intro rule with none*)  | 
|
| 42360 | 311  | 
if is_nontrivial (Proof_Context.theory_of ctxt) (major_prem_of theorem)  | 
| 17205 | 312  | 
andalso not (null successful)  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
313  | 
then SOME (nprems_of theorem - 1, foldl1 Int.min successful) else NONE  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
314  | 
end  | 
| 46718 | 315  | 
else NONE;  | 
| 16036 | 316  | 
|
| 32738 | 317  | 
val tac_limit = Unsynchronized.ref 5;  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
318  | 
|
| 30143 | 319  | 
fun filter_solves ctxt goal =  | 
320  | 
let  | 
|
321  | 
fun etacn thm i = Seq.take (! tac_limit) o etac thm i;  | 
|
322  | 
fun try_thm thm =  | 
|
323  | 
if Thm.no_prems thm then rtac thm 1 goal  | 
|
| 
30318
 
3d03190d2864
replaced archaic use of rep_ss by Simplifier.mksimps;
 
wenzelm 
parents: 
30234 
diff
changeset
 | 
324  | 
else (etacn thm THEN_ALL_NEW (Goal.norm_hhf_tac THEN' Method.assm_tac ctxt)) 1 goal;  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
325  | 
in  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
326  | 
fn Internal (_, thm) =>  | 
| 43620 | 327  | 
if is_some (Seq.pull (try_thm thm))  | 
328  | 
then SOME (Thm.nprems_of thm, 0) else NONE  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
329  | 
| External _ => NONE  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
330  | 
end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
331  | 
|
| 
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
 | 
332  | 
|
| 16074 | 333  | 
(* filter_simp *)  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
335  | 
fun filter_simp ctxt t (Internal (_, thm)) =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
336  | 
let  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
337  | 
val mksimps = Simplifier.mksimps (simpset_of ctxt);  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
338  | 
val extract_simp =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
339  | 
(map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl);  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
340  | 
val ss = is_matching_thm extract_simp ctxt false t thm;  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
341  | 
in  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
342  | 
if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
343  | 
end  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
344  | 
| filter_simp _ _ (External _) = NONE;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
345  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
347  | 
(* filter_pattern *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
|
| 32798 | 349  | 
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
 | 
350  | 
|
| 30143 | 351  | 
(*Including all constants and frees is only sound because  | 
352  | 
matching uses higher-order patterns. If full matching  | 
|
353  | 
were used, then constants that may be subject to  | 
|
354  | 
beta-reduction after substitution of frees should  | 
|
355  | 
not be included for LHS set because they could be  | 
|
356  | 
thrown away by the substituted function.  | 
|
357  | 
e.g. for (?F 1 2) do not include 1 or 2, if it were  | 
|
358  | 
possible for ?F to be (% x y. 3)  | 
|
359  | 
The largest possible set should always be included on  | 
|
360  | 
the RHS.*)  | 
|
361  | 
||
362  | 
fun filter_pattern ctxt pat =  | 
|
363  | 
let  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
364  | 
val pat_consts = get_names pat;  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
365  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
366  | 
fun check (theorem, NONE) = check (theorem, SOME (get_names (prop_of theorem)))  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
367  | 
| check (theorem, c as SOME thm_consts) =  | 
| 33038 | 368  | 
(if subset (op =) (pat_consts, thm_consts) andalso  | 
| 42360 | 369  | 
Pattern.matches_subterm (Proof_Context.theory_of ctxt) (pat, prop_of theorem)  | 
| 32798 | 370  | 
then SOME (0, 0) else NONE, c);  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
371  | 
in check end;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
|
| 
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
 | 
373  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
374  | 
(* interpret criteria as filters *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
375  | 
|
| 16036 | 376  | 
local  | 
377  | 
||
378  | 
fun err_no_goal c =  | 
|
379  | 
  error ("Current goal required for " ^ c ^ " search criterion");
 | 
|
380  | 
||
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
381  | 
val fix_goal = Thm.prop_of;  | 
| 
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
382  | 
|
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
383  | 
fun filter_crit _ _ (Name name) = apfst (filter_name name)  | 
| 16036 | 384  | 
| filter_crit _ NONE Intro = err_no_goal "intro"  | 
385  | 
| filter_crit _ NONE Elim = err_no_goal "elim"  | 
|
386  | 
| 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
 | 
387  | 
| filter_crit _ NONE Solves = err_no_goal "solves"  | 
| 
46717
 
b09afce1e54f
removed broken/unused introiff (cf. d452117ba564);
 
wenzelm 
parents: 
46716 
diff
changeset
 | 
388  | 
| filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt (fix_goal goal))  | 
| 30143 | 389  | 
| filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (fix_goal goal))  | 
390  | 
| filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (fix_goal goal))  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
391  | 
| filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal)  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
392  | 
| filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat)  | 
| 16088 | 393  | 
| filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;  | 
| 16036 | 394  | 
|
| 19502 | 395  | 
fun opt_not x = if is_some x then NONE else SOME (0, 0);  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
396  | 
|
| 17756 | 397  | 
fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int)  | 
| 26283 | 398  | 
| opt_add _ _ = NONE;  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
399  | 
|
| 30143 | 400  | 
fun app_filters thm =  | 
401  | 
let  | 
|
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
402  | 
fun app (NONE, _, _) = NONE  | 
| 32798 | 403  | 
| app (SOME v, _, []) = SOME (v, thm)  | 
| 30143 | 404  | 
| app (r, consts, f :: fs) =  | 
405  | 
let val (r', consts') = f (thm, consts)  | 
|
406  | 
in app (opt_add r r', consts', fs) end;  | 
|
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
407  | 
in app end;  | 
| 
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
408  | 
|
| 16036 | 409  | 
in  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
410  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
411  | 
fun filter_criterion ctxt opt_goal (b, c) =  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
412  | 
(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
 | 
413  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
414  | 
fun sorted_filter filters theorems =  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
415  | 
let  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
416  | 
fun eval_filters theorem = app_filters theorem (SOME (0, 0), NONE, filters);  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
417  | 
|
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
418  | 
(*filters return: (number of assumptions, substitution size) option, so  | 
| 
16964
 
6a25e42eaff5
Ordering is now: first by number of assumptions, second by the substitution size.
 
kleing 
parents: 
16895 
diff
changeset
 | 
419  | 
sort (desc. in both cases) according to number of assumptions first,  | 
| 
16895
 
df67fc190e06
Sort search results in order of relevance, where relevance =
 
kleing 
parents: 
16486 
diff
changeset
 | 
420  | 
then by the substitution size*)  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
421  | 
fun result_ord (((p0, s0), _), ((p1, s1), _)) =  | 
| 17205 | 422  | 
prod_ord int_ord int_ord ((p1, s1), (p0, s0));  | 
| 46977 | 423  | 
in  | 
424  | 
grouped 100 Par_List.map eval_filters theorems  | 
|
425  | 
|> map_filter I |> sort result_ord |> map #2  | 
|
426  | 
end;  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
427  | 
|
| 30822 | 428  | 
fun lazy_filter filters =  | 
429  | 
let  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
430  | 
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
 | 
431  | 
and first_match [] = NONE  | 
| 30822 | 432  | 
| first_match (thm :: thms) =  | 
433  | 
(case app_filters thm (SOME (0, 0), NONE, filters) of  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
434  | 
NONE => first_match thms  | 
| 30822 | 435  | 
| SOME (_, t) => SOME (t, lazy_match thms));  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
436  | 
in lazy_match end;  | 
| 30822 | 437  | 
|
| 16036 | 438  | 
end;  | 
439  | 
||
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
440  | 
|
| 22414 | 441  | 
(* removing duplicates, preferring nicer names, roughly n log n *)  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
442  | 
|
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
443  | 
local  | 
| 
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
444  | 
|
| 27486 | 445  | 
val index_ord = option_ord (K EQUAL);  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33039 
diff
changeset
 | 
446  | 
val hidden_ord = bool_ord o pairself Name_Space.is_hidden;  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30318 
diff
changeset
 | 
447  | 
val qual_ord = int_ord o pairself (length o Long_Name.explode);  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
448  | 
val txt_ord = int_ord o pairself size;  | 
| 
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
449  | 
|
| 27486 | 450  | 
fun nicer_name (x, i) (y, j) =  | 
451  | 
(case hidden_ord (x, y) of EQUAL =>  | 
|
452  | 
(case index_ord (i, j) of EQUAL =>  | 
|
453  | 
(case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord)  | 
|
454  | 
| ord => ord)  | 
|
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
455  | 
| ord => ord) <> GREATER;  | 
| 
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
456  | 
|
| 29848 | 457  | 
fun rem_cdups nicer xs =  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
458  | 
let  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
459  | 
fun rem_c rev_seen [] = rev rev_seen  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
460  | 
| rem_c rev_seen [x] = rem_c (x :: rev_seen) []  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
461  | 
| rem_c rev_seen ((x as (t, _)) :: (y as (t', _)) :: xs) =  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
462  | 
if (prop_of t) aconv (prop_of t')  | 
| 
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
463  | 
then rem_c rev_seen ((if nicer (fact_ref_of t) (fact_ref_of t') then x else y) :: xs)  | 
| 30822 | 464  | 
else rem_c (x :: rev_seen) (y :: xs)  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
465  | 
in rem_c [] xs end;  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
466  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
467  | 
in  | 
| 
25226
 
502d8676cdd6
improved notion of 'nicer' fact names (observe some name space properties);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
468  | 
|
| 30143 | 469  | 
fun nicer_shortest ctxt =  | 
470  | 
let  | 
|
| 46718 | 471  | 
(* FIXME Why global name space!?? *)  | 
| 42360 | 472  | 
val space = Facts.space_of (Global_Theory.facts_of (Proof_Context.theory_of ctxt));  | 
| 29848 | 473  | 
|
| 
30216
 
0300b7420b07
nicer_shortest: use NameSpace.extern_flags with disabled "features" instead of internal NameSpace.get_accesses;
 
wenzelm 
parents: 
30188 
diff
changeset
 | 
474  | 
val shorten =  | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
475  | 
Name_Space.extern  | 
| 
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
476  | 
(ctxt  | 
| 
42669
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
477  | 
|> Config.put Name_Space.names_long false  | 
| 
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
478  | 
|> Config.put Name_Space.names_short false  | 
| 
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
479  | 
|> Config.put Name_Space.names_unique false) space;  | 
| 29848 | 480  | 
|
481  | 
fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =  | 
|
482  | 
nicer_name (shorten x, i) (shorten y, j)  | 
|
483  | 
| nicer (Facts.Fact _) (Facts.Named _) = true  | 
|
484  | 
| nicer (Facts.Named _) (Facts.Fact _) = false;  | 
|
485  | 
in nicer end;  | 
|
486  | 
||
487  | 
fun rem_thm_dups nicer xs =  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
488  | 
xs ~~ (1 upto length xs)  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
489  | 
|> sort (Term_Ord.fast_term_ord o pairself (prop_of o #1))  | 
| 29848 | 490  | 
|> rem_cdups nicer  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
491  | 
|> sort (int_ord o pairself #2)  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
492  | 
|> map #1;  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
493  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
26283 
diff
changeset
 | 
494  | 
end;  | 
| 
22340
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
495  | 
|
| 
 
275802767bf3
Remove duplicates from printed theorems in find_theorems
 
kleing 
parents: 
19502 
diff
changeset
 | 
496  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
497  | 
(* print_theorems *)  | 
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
498  | 
|
| 26283 | 499  | 
fun all_facts_of ctxt =  | 
| 33381 | 500  | 
let  | 
| 33382 | 501  | 
fun visible_facts facts =  | 
502  | 
Facts.dest_static [] facts  | 
|
503  | 
|> filter_out (Facts.is_concealed facts o #1);  | 
|
| 33381 | 504  | 
in  | 
505  | 
maps Facts.selections  | 
|
| 42360 | 506  | 
(visible_facts (Global_Theory.facts_of (Proof_Context.theory_of ctxt)) @  | 
507  | 
visible_facts (Proof_Context.facts_of ctxt))  | 
|
| 33381 | 508  | 
end;  | 
| 17972 | 509  | 
|
| 32738 | 510  | 
val limit = Unsynchronized.ref 40;  | 
| 
25992
 
928594f50893
renamed thms_containing_limit to FindTheorems.limit;
 
wenzelm 
parents: 
25226 
diff
changeset
 | 
511  | 
|
| 43070 | 512  | 
fun filter_theorems ctxt theorems query =  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
513  | 
let  | 
| 46718 | 514  | 
    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
 | 
515  | 
val filters = map (filter_criterion ctxt opt_goal) criteria;  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
516  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
517  | 
fun find_all theorems =  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
518  | 
let  | 
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
519  | 
val raw_matches = sorted_filter filters theorems;  | 
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
520  | 
|
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
521  | 
val matches =  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
522  | 
if rem_dups  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
523  | 
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
 | 
524  | 
else raw_matches;  | 
| 
28900
 
53fd5cc685b4
FindTheorems performance improvements (from Timothy Bourke)
 
kleing 
parents: 
28211 
diff
changeset
 | 
525  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
526  | 
val len = length matches;  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
527  | 
val lim = the_default (! limit) opt_limit;  | 
| 34088 | 528  | 
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
 | 
529  | 
|
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
530  | 
val find =  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
531  | 
if rem_dups orelse is_none opt_limit  | 
| 
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
532  | 
then find_all  | 
| 30822 | 533  | 
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
 | 
534  | 
|
| 
41844
 
b933142e02d0
generalize find_theorems filters to work on raw propositions, too
 
krauss 
parents: 
41841 
diff
changeset
 | 
535  | 
in find theorems end;  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
536  | 
|
| 46718 | 537  | 
fun filter_theorems_cmd ctxt theorems raw_query =  | 
538  | 
filter_theorems ctxt theorems (map_criteria  | 
|
| 43070 | 539  | 
(map (apsnd (read_criterion ctxt))) raw_query);  | 
| 43067 | 540  | 
|
541  | 
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
 | 
542  | 
let  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
543  | 
val assms =  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
544  | 
Proof_Context.get_fact ctxt (Facts.named "local.assms")  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
545  | 
handle ERROR _ => [];  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
546  | 
val add_prems = Seq.hd o TRY (Method.insert_tac assms 1);  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
547  | 
val opt_goal' = Option.map add_prems opt_goal;  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
548  | 
in  | 
| 46718 | 549  | 
filter ctxt (map Internal (all_facts_of ctxt))  | 
550  | 
      {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
 | 
551  | 
|> apsnd (map (fn Internal f => f))  | 
| 
 
88e45168272c
moved questionable goal modification out of filter_theorems
 
krauss 
parents: 
43068 
diff
changeset
 | 
552  | 
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
 | 
553  | 
|
| 43067 | 554  | 
val find_theorems = gen_find_theorems filter_theorems;  | 
555  | 
val find_theorems_cmd = gen_find_theorems filter_theorems_cmd;  | 
|
556  | 
||
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
557  | 
fun pretty_theorem ctxt (Internal (thmref, thm)) = Pretty.block  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
558  | 
[Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1,  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
559  | 
Display.pretty_thm ctxt thm]  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
560  | 
| pretty_theorem ctxt (External (thmref, prop)) = Pretty.block  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
561  | 
[Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1,  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
562  | 
Syntax.unparse_term ctxt prop];  | 
| 
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
 | 
563  | 
|
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
564  | 
fun pretty_thm ctxt (thmref, thm) = pretty_theorem ctxt (Internal (thmref, thm));  | 
| 
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
565  | 
|
| 
43076
 
7b06cd71792c
parameterize print_theorems over actual search function
 
krauss 
parents: 
43071 
diff
changeset
 | 
566  | 
fun gen_print_theorems find ctxt opt_goal opt_limit rem_dups raw_criteria =  | 
| 30143 | 567  | 
let  | 
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29848 
diff
changeset
 | 
568  | 
val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;  | 
| 
43076
 
7b06cd71792c
parameterize print_theorems over actual search function
 
krauss 
parents: 
43071 
diff
changeset
 | 
569  | 
val (foundo, theorems) = find  | 
| 46718 | 570  | 
      {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
 | 
571  | 
val returned = length theorems;  | 
| 
31684
 
d5d830979a54
minor tuning according to Isabelle/ML conventions;
 
wenzelm 
parents: 
31042 
diff
changeset
 | 
572  | 
|
| 
30785
 
15f64e05e703
Limit the number of results returned by auto_solves.
 
Timothy Bourke 
parents: 
30693 
diff
changeset
 | 
573  | 
val tally_msg =  | 
| 30822 | 574  | 
(case foundo of  | 
| 38335 | 575  | 
NONE => "displaying " ^ string_of_int returned ^ " theorem(s)"  | 
| 30822 | 576  | 
| SOME found =>  | 
| 38335 | 577  | 
"found " ^ string_of_int found ^ " theorem(s)" ^  | 
| 30822 | 578  | 
(if returned < found  | 
579  | 
             then " (" ^ string_of_int returned ^ " displayed)"
 | 
|
580  | 
else ""));  | 
|
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
581  | 
in  | 
| 38335 | 582  | 
Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) ::  | 
583  | 
Pretty.str "" ::  | 
|
| 
46716
 
c45a4427db39
discontinued slightly odd built-in timing (cf. 53fd5cc685b4) -- the Isar toplevel does that already (e.g. via Toplevel.timing or Toplevel.profiling);
 
wenzelm 
parents: 
46713 
diff
changeset
 | 
584  | 
(if null theorems then [Pretty.str "nothing found"]  | 
| 38335 | 585  | 
else  | 
| 
46716
 
c45a4427db39
discontinued slightly odd built-in timing (cf. 53fd5cc685b4) -- the Isar toplevel does that already (e.g. via Toplevel.timing or Toplevel.profiling);
 
wenzelm 
parents: 
46713 
diff
changeset
 | 
586  | 
[Pretty.str (tally_msg ^ ":"), Pretty.str ""] @  | 
| 46977 | 587  | 
grouped 10 Par_List.map (pretty_theorem ctxt) theorems)  | 
| 38335 | 588  | 
end |> Pretty.chunks |> Pretty.writeln;  | 
| 
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
 | 
589  | 
|
| 
43076
 
7b06cd71792c
parameterize print_theorems over actual search function
 
krauss 
parents: 
43071 
diff
changeset
 | 
590  | 
fun print_theorems ctxt =  | 
| 
 
7b06cd71792c
parameterize print_theorems over actual search function
 
krauss 
parents: 
43071 
diff
changeset
 | 
591  | 
gen_print_theorems (filter_theorems ctxt (map Internal (all_facts_of ctxt))) ctxt;  | 
| 
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
 | 
592  | 
|
| 32798 | 593  | 
|
| 46718 | 594  | 
|
| 
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
 | 
595  | 
(** command syntax **)  | 
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
596  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
597  | 
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
 | 
598  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
599  | 
val criterion =  | 
| 36950 | 600  | 
Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name ||  | 
601  | 
Parse.reserved "intro" >> K Intro ||  | 
|
602  | 
Parse.reserved "elim" >> K Elim ||  | 
|
603  | 
Parse.reserved "dest" >> K Dest ||  | 
|
604  | 
Parse.reserved "solves" >> K Solves ||  | 
|
605  | 
Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp ||  | 
|
606  | 
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
 | 
607  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
608  | 
val options =  | 
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
609  | 
Scan.optional  | 
| 36950 | 610  | 
    (Parse.$$$ "(" |--
 | 
611  | 
Parse.!!! (Scan.option Parse.nat -- Scan.optional (Parse.reserved "with_dups" >> K false) true  | 
|
612  | 
--| Parse.$$$ ")")) (NONE, true);  | 
|
| 
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
 | 
613  | 
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
 | 
614  | 
|
| 
43068
 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 
krauss 
parents: 
43067 
diff
changeset
 | 
615  | 
val query_parser = Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion));  | 
| 
 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 
krauss 
parents: 
43067 
diff
changeset
 | 
616  | 
|
| 
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
 | 
617  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46718 
diff
changeset
 | 
618  | 
  Outer_Syntax.improper_command ("find_theorems", Keyword.diag)
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46718 
diff
changeset
 | 
619  | 
"print theorems meeting specified criteria"  | 
| 
43068
 
ac769b5edd1d
exported raw query parser; removed inconsistent clone
 
krauss 
parents: 
43067 
diff
changeset
 | 
620  | 
(options -- query_parser  | 
| 38334 | 621  | 
>> (fn ((opt_lim, rem_dups), spec) =>  | 
622  | 
Toplevel.no_timing o  | 
|
623  | 
Toplevel.keep (fn state =>  | 
|
624  | 
let  | 
|
625  | 
val ctxt = Toplevel.context_of state;  | 
|
626  | 
val opt_goal = try (Proof.simple_goal o Toplevel.proof_of) state |> Option.map #goal;  | 
|
| 
41845
 
6611b9cef38b
reactivate time measurement (partly reverting c27b0b37041a);
 
krauss 
parents: 
41844 
diff
changeset
 | 
627  | 
in print_theorems ctxt opt_goal opt_lim rem_dups spec end)));  | 
| 
16033
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
628  | 
|
| 
 
f93ca3d4ffa7
Retrieve theorems from proof context -- improved version of
 
wenzelm 
parents:  
diff
changeset
 | 
629  | 
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
 | 
630  | 
|
| 
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
631  | 
end;  |