author | blanchet |
Wed, 14 Jul 2021 16:09:57 +0200 | |
changeset 73980 | 291f7b5fc7c9 |
parent 73979 | e5322146e7e8 |
child 74379 | 9ea507f63bcb |
permissions | -rw-r--r-- |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact.ML |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
2 |
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
4 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
5 |
Sledgehammer fact handling. |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
6 |
*) |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
7 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
8 |
signature SLEDGEHAMMER_FACT = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
9 |
sig |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
10 |
type status = ATP_Problem_Generate.status |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
11 |
type stature = ATP_Problem_Generate.stature |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
12 |
|
73979 | 13 |
type lazy_fact = ((unit -> string) * stature) * thm |
51004
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents:
50815
diff
changeset
|
14 |
type fact = (string * stature) * thm |
48296
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48292
diff
changeset
|
15 |
|
48292 | 16 |
type fact_override = |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
17 |
{add : (Facts.ref * Token.src list) list, |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
18 |
del : (Facts.ref * Token.src list) list, |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
19 |
only : bool} |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
20 |
|
48292 | 21 |
val no_fact_override : fact_override |
55212 | 22 |
|
58919 | 23 |
val fact_of_ref : Proof.context -> Keyword.keywords -> thm list -> status Termtab.table -> |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
24 |
Facts.ref * Token.src list -> ((string * stature) * thm) list |
62095 | 25 |
val cartouche_thm : Proof.context -> thm -> string |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
26 |
val is_blacklisted_or_something : string -> bool |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
27 |
val clasimpset_rule_table_of : Proof.context -> status Termtab.table |
55212 | 28 |
val build_name_tables : (thm -> string) -> ('a * thm) list -> |
29 |
string Symtab.table * string Symtab.table |
|
61322 | 30 |
val fact_distinct : (term * term -> bool) -> ('a * thm) list -> ('a * thm) list |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
31 |
val instantiate_inducts : Proof.context -> term list -> term -> |
55212 | 32 |
(((unit -> string) * 'a) * thm) list -> (((unit -> string) * 'a) * thm) list |
73979 | 33 |
val fact_of_lazy_fact : lazy_fact -> fact |
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
34 |
val is_useful_unnamed_local_fact : Proof.context -> thm -> bool |
55212 | 35 |
|
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
36 |
val all_facts : Proof.context -> bool -> Keyword.keywords -> thm list -> thm list -> |
73979 | 37 |
status Termtab.table -> lazy_fact list |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
38 |
val nearly_all_facts : Proof.context -> bool -> fact_override -> Keyword.keywords -> |
73979 | 39 |
status Termtab.table -> thm list -> term list -> term -> lazy_fact list |
40 |
val drop_duplicate_facts : lazy_fact list -> lazy_fact list |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
41 |
end; |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
42 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
43 |
structure Sledgehammer_Fact : SLEDGEHAMMER_FACT = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
44 |
struct |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
45 |
|
50495
bd9a0028b063
better tautology check -- don't reject "prod_cases3" for example
blanchet
parents:
50490
diff
changeset
|
46 |
open ATP_Util |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
47 |
open ATP_Problem_Generate |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
48 |
open Sledgehammer_Util |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
49 |
|
73979 | 50 |
type lazy_fact = ((unit -> string) * stature) * thm |
51004
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents:
50815
diff
changeset
|
51 |
type fact = (string * stature) * thm |
48296
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48292
diff
changeset
|
52 |
|
48292 | 53 |
type fact_override = |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
54 |
{add : (Facts.ref * Token.src list) list, |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
55 |
del : (Facts.ref * Token.src list) list, |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
56 |
only : bool} |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
57 |
|
60252
2c468c062589
improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
blanchet
parents:
59970
diff
changeset
|
58 |
val local_thisN = Long_Name.localN ^ Long_Name.separator ^ Auto_Bind.thisN |
2c468c062589
improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
blanchet
parents:
59970
diff
changeset
|
59 |
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
60 |
(* gracefully handle huge background theories *) |
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
61 |
val max_facts_for_duplicates = 50000 |
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
62 |
val max_facts_for_complex_check = 25000 |
54083
824db6ab3339
crank up limit a bit -- truly huge background theories are still nearly 3 times larger
blanchet
parents:
54081
diff
changeset
|
63 |
val max_simps_for_clasimpset = 10000 |
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
64 |
|
48292 | 65 |
val no_fact_override = {add = [], del = [], only = false} |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
66 |
|
58919 | 67 |
fun needs_quoting keywords s = |
58928
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58919
diff
changeset
|
68 |
Keyword.is_literal keywords s orelse |
50239 | 69 |
exists (not o Symbol_Pos.is_identifier) (Long_Name.explode s) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
70 |
|
58919 | 71 |
fun make_name keywords multi j name = |
72 |
(name |> needs_quoting keywords name ? quote) ^ |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
73 |
(if multi then "(" ^ string_of_int j ^ ")" else "") |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
74 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
75 |
fun explode_interval _ (Facts.FromTo (i, j)) = i upto j |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
76 |
| explode_interval max (Facts.From i) = i upto i + max - 1 |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
77 |
| explode_interval _ (Facts.Single i) = [i] |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
78 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
79 |
fun is_rec_eq lhs = Term.exists_subterm (curry (op =) (head_of lhs)) |
57468 | 80 |
|
69597 | 81 |
fun is_rec_def (\<^const>\<open>Trueprop\<close> $ t) = is_rec_def t |
82 |
| is_rec_def (\<^const>\<open>Pure.imp\<close> $ _ $ t2) = is_rec_def t2 |
|
69593 | 83 |
| is_rec_def (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t1 $ t2) = is_rec_eq t1 t2 |
84 |
| is_rec_def (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ t2) = is_rec_eq t1 t2 |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
85 |
| is_rec_def _ = false |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
86 |
|
59582 | 87 |
fun is_assum assms th = exists (fn ct => Thm.prop_of th aconv Thm.term_of ct) assms |
48396 | 88 |
fun is_chained chained = member Thm.eq_thm_prop chained |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
89 |
|
48396 | 90 |
fun scope_of_thm global assms chained th = |
91 |
if is_chained chained th then Chained |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
92 |
else if global then Global |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
93 |
else if is_assum assms th then Assum |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
94 |
else Local |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
95 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
96 |
val may_be_induction = |
69593 | 97 |
exists_subterm (fn Var (_, Type (\<^type_name>\<open>fun\<close>, [_, T])) => body_type T = \<^typ>\<open>bool\<close> |
57729 | 98 |
| _ => false) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
99 |
|
54077 | 100 |
(* TODO: get rid of *) |
53501 | 101 |
fun normalize_vars t = |
102 |
let |
|
103 |
fun normT (Type (s, Ts)) = fold_map normT Ts #>> curry Type s |
|
104 |
| normT (TVar (z as (_, S))) = |
|
105 |
(fn ((knownT, nT), accum) => |
|
55286 | 106 |
(case find_index (equal z) knownT of |
53501 | 107 |
~1 => (TVar ((Name.uu, nT), S), ((z :: knownT, nT + 1), accum)) |
55286 | 108 |
| j => (TVar ((Name.uu, nT - j - 1), S), ((knownT, nT), accum)))) |
53501 | 109 |
| normT (T as TFree _) = pair T |
57468 | 110 |
|
53501 | 111 |
fun norm (t $ u) = norm t ##>> norm u #>> op $ |
112 |
| norm (Const (s, T)) = normT T #>> curry Const s |
|
57729 | 113 |
| norm (Var (z as (_, T))) = normT T |
53501 | 114 |
#> (fn (T, (accumT, (known, n))) => |
57729 | 115 |
(case find_index (equal z) known of |
116 |
~1 => (Var ((Name.uu, n), T), (accumT, (z :: known, n + 1))) |
|
117 |
| j => (Var ((Name.uu, n - j - 1), T), (accumT, (known, n))))) |
|
118 |
| norm (Abs (_, T, t)) = norm t ##>> normT T #>> (fn (t, T) => Abs (Name.uu, T, t)) |
|
53501 | 119 |
| norm (Bound j) = pair (Bound j) |
120 |
| norm (Free (s, T)) = normT T #>> curry Free s |
|
121 |
in fst (norm t (([], 0), ([], 0))) end |
|
122 |
||
48396 | 123 |
fun status_of_thm css name th = |
54081 | 124 |
if Termtab.is_empty css then |
125 |
General |
|
126 |
else |
|
59582 | 127 |
let val t = Thm.prop_of th in |
54081 | 128 |
(* FIXME: use structured name *) |
129 |
if String.isSubstring ".induct" name andalso may_be_induction t then |
|
130 |
Induction |
|
131 |
else |
|
132 |
let val t = normalize_vars t in |
|
55286 | 133 |
(case Termtab.lookup css t of |
54081 | 134 |
SOME status => status |
135 |
| NONE => |
|
136 |
let val concl = Logic.strip_imp_concl t in |
|
55286 | 137 |
(case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of |
54081 | 138 |
SOME lrhss => |
139 |
let |
|
140 |
val prems = Logic.strip_imp_prems t |
|
141 |
val t' = Logic.list_implies (prems, Logic.mk_equals lrhss) |
|
142 |
in |
|
143 |
Termtab.lookup css t' |> the_default General |
|
144 |
end |
|
55286 | 145 |
| NONE => General) |
146 |
end) |
|
54081 | 147 |
end |
148 |
end |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
149 |
|
48396 | 150 |
fun stature_of_thm global assms chained css name th = |
151 |
(scope_of_thm global assms chained th, status_of_thm css name th) |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
152 |
|
58919 | 153 |
fun fact_of_ref ctxt keywords chained css (xthm as (xref, args)) = |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
154 |
let |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
155 |
val ths = Attrib.eval_thms ctxt [xthm] |
61877
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
wenzelm
parents:
61646
diff
changeset
|
156 |
val bracket = |
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
wenzelm
parents:
61646
diff
changeset
|
157 |
implode (map (enclose "[" "]" o Pretty.unformatted_string_of o Token.pretty_src ctxt) args) |
57468 | 158 |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
159 |
fun nth_name j = |
55286 | 160 |
(case xref of |
62095 | 161 |
Facts.Fact s => cartouche (simplify_spaces (YXML.content_of s)) ^ bracket |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
162 |
| Facts.Named (("", _), _) => "[" ^ bracket ^ "]" |
58919 | 163 |
| Facts.Named ((name, _), NONE) => make_name keywords (length ths > 1) (j + 1) name ^ bracket |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
164 |
| Facts.Named ((name, _), SOME intervals) => |
58919 | 165 |
make_name keywords true |
55286 | 166 |
(nth (maps (explode_interval (length ths)) intervals) j) name ^ bracket) |
57468 | 167 |
|
48396 | 168 |
fun add_nth th (j, rest) = |
169 |
let val name = nth_name j in |
|
57729 | 170 |
(j + 1, ((name, stature_of_thm false [] chained css name th), th) :: rest) |
48396 | 171 |
end |
57468 | 172 |
in |
173 |
(0, []) |> fold add_nth ths |> snd |
|
174 |
end |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
175 |
|
57729 | 176 |
(* Reject theorems with names like "List.filter.filter_list_def" or "Accessible_Part.acc.defs", as |
177 |
these are definitions arising from packages. *) |
|
50736 | 178 |
fun is_package_def s = |
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
179 |
exists (fn suf => String.isSuffix suf s) |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
180 |
["_case_def", "_rec_def", "_size_def", "_size_overloaded_def"] |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
181 |
andalso |
50736 | 182 |
let val ss = Long_Name.explode s in |
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
183 |
length ss > 2 andalso not (hd ss = "local") |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
184 |
end |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
185 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
186 |
(* FIXME: put other record thms here, or declare as "no_atp" *) |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
187 |
val multi_base_blacklist = |
57729 | 188 |
["defs", "select_defs", "update_defs", "split", "splits", "split_asm", "ext_cases", "eq.simps", |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57963
diff
changeset
|
189 |
"eq.refl", "nchotomy", "case_cong", "case_cong_weak", "nat_of_char_simps", "nibble.simps", |
48440
159e320ef851
identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents:
48438
diff
changeset
|
190 |
"nibble.distinct"] |
159e320ef851
identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents:
48438
diff
changeset
|
191 |
|> map (prefix Long_Name.separator) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
192 |
|
57729 | 193 |
(* The maximum apply depth of any "metis" call in "Metis_Examples" (back in 2007) was 11. *) |
53529
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
194 |
val max_apply_depth = 18 |
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
195 |
|
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
196 |
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1) |
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
197 |
| apply_depth (Abs (_, _, t)) = apply_depth t |
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
198 |
| apply_depth _ = 0 |
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
199 |
|
53533 | 200 |
fun is_too_complex t = apply_depth t > max_apply_depth |
53529
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
201 |
|
48667
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
202 |
(* FIXME: Ad hoc list *) |
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
203 |
val technical_prefixes = |
57729 | 204 |
["ATP", "Code_Evaluation", "Datatype", "Enum", "Lazy_Sequence", "Limited_Sequence", "Meson", |
205 |
"Metis", "Nitpick", "Quickcheck_Random", "Quickcheck_Exhaustive", "Quickcheck_Narrowing", |
|
50736 | 206 |
"Random_Sequence", "Sledgehammer", "SMT"] |
48667
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
207 |
|> map (suffix Long_Name.separator) |
48440
159e320ef851
identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents:
48438
diff
changeset
|
208 |
|
57729 | 209 |
fun is_technical_const s = exists (fn pref => String.isPrefix pref s) technical_prefixes |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
210 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
211 |
(* FIXME: make more reliable *) |
53507 | 212 |
val sep_class_sep = Long_Name.separator ^ "class" ^ Long_Name.separator |
57468 | 213 |
|
57729 | 214 |
fun is_low_level_class_const s = |
69593 | 215 |
s = \<^const_name>\<open>equal_class.equal\<close> orelse String.isSubstring sep_class_sep s |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
216 |
|
60448 | 217 |
val sep_that = Long_Name.separator ^ Auto_Bind.thatN |
54081 | 218 |
val skolem_thesis = Name.skolem Auto_Bind.thesisN |
219 |
||
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
220 |
fun is_that_fact th = |
59582 | 221 |
exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (Thm.prop_of th) |
54081 | 222 |
andalso String.isSuffix sep_that (Thm.get_name_hint th) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
223 |
|
53513 | 224 |
datatype interest = Deal_Breaker | Interesting | Boring |
225 |
||
226 |
fun combine_interests Deal_Breaker _ = Deal_Breaker |
|
227 |
| combine_interests _ Deal_Breaker = Deal_Breaker |
|
228 |
| combine_interests Interesting _ = Interesting |
|
229 |
| combine_interests _ Interesting = Interesting |
|
230 |
| combine_interests Boring Boring = Boring |
|
231 |
||
54914
25212efcd0de
removed pointless warning (cf. http://stackoverflow.com/questions/20233463/isabelle-metis-proof-state-contains-the-universal-sort/20785045#20785045)
blanchet
parents:
54503
diff
changeset
|
232 |
val type_has_top_sort = |
25212efcd0de
removed pointless warning (cf. http://stackoverflow.com/questions/20233463/isabelle-metis-proof-state-contains-the-universal-sort/20785045#20785045)
blanchet
parents:
54503
diff
changeset
|
233 |
exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false) |
25212efcd0de
removed pointless warning (cf. http://stackoverflow.com/questions/20233463/isabelle-metis-proof-state-contains-the-universal-sort/20785045#20785045)
blanchet
parents:
54503
diff
changeset
|
234 |
|
50523
0799339fea0f
get rid of some junk facts in the MaSh evaluation driver
blanchet
parents:
50512
diff
changeset
|
235 |
fun is_likely_tautology_too_meta_or_too_technical th = |
48406 | 236 |
let |
50053
fea589c8583e
fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents:
50049
diff
changeset
|
237 |
fun is_interesting_subterm (Const (s, _)) = |
fea589c8583e
fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents:
50049
diff
changeset
|
238 |
not (member (op =) atp_widely_irrelevant_consts s) |
fea589c8583e
fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents:
50049
diff
changeset
|
239 |
| is_interesting_subterm (Free _) = true |
fea589c8583e
fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents:
50049
diff
changeset
|
240 |
| is_interesting_subterm _ = false |
57468 | 241 |
|
53513 | 242 |
fun interest_of_bool t = |
57729 | 243 |
if exists_Const ((is_technical_const o fst) orf (is_low_level_class_const o fst) orf |
244 |
type_has_top_sort o snd) t then |
|
53513 | 245 |
Deal_Breaker |
69593 | 246 |
else if exists_type (exists_subtype (curry (op =) \<^typ>\<open>prop\<close>)) t orelse |
57729 | 247 |
not (exists_subterm is_interesting_subterm t) then |
53513 | 248 |
Boring |
249 |
else |
|
250 |
Interesting |
|
57468 | 251 |
|
69597 | 252 |
fun interest_of_prop _ (\<^const>\<open>Trueprop\<close> $ t) = interest_of_bool t |
253 |
| interest_of_prop Ts (\<^const>\<open>Pure.imp\<close> $ t $ u) = |
|
53513 | 254 |
combine_interests (interest_of_prop Ts t) (interest_of_prop Ts u) |
69593 | 255 |
| interest_of_prop Ts (Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs (_, T, t)) = |
54040 | 256 |
if type_has_top_sort T then Deal_Breaker else interest_of_prop (T :: Ts) t |
69593 | 257 |
| interest_of_prop Ts ((t as Const (\<^const_name>\<open>Pure.all\<close>, _)) $ u) = |
53513 | 258 |
interest_of_prop Ts (t $ eta_expand Ts u 1) |
69593 | 259 |
| interest_of_prop _ (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t $ u) = |
53513 | 260 |
combine_interests (interest_of_bool t) (interest_of_bool u) |
261 |
| interest_of_prop _ _ = Deal_Breaker |
|
57468 | 262 |
|
59582 | 263 |
val t = Thm.prop_of th |
48406 | 264 |
in |
57729 | 265 |
(interest_of_prop [] t <> Interesting andalso not (Thm.eq_thm_prop (@{thm ext}, th))) orelse |
53507 | 266 |
is_that_fact th |
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48437
diff
changeset
|
267 |
end |
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48437
diff
changeset
|
268 |
|
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
269 |
val is_blacklisted_or_something = |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
270 |
let val blist = multi_base_blacklist in |
57468 | 271 |
fn name => is_package_def name orelse exists (fn s => String.isSuffix s name) blist |
272 |
end |
|
50510 | 273 |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
274 |
(* This is a terrible hack. Free variables are sometimes coded as "M__" when |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
275 |
they are displayed as "M" and we want to avoid clashes with these. But |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
276 |
sometimes it's even worse: "Ma__" encodes "M". So we simply reserve all |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
277 |
prefixes of all free variables. In the worse case scenario, where the fact |
49981 | 278 |
won't be resolved correctly, the user can fix it manually, e.g., by giving a |
279 |
name to the offending fact. *) |
|
57729 | 280 |
fun all_prefixes_of s = map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
281 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
282 |
fun close_form t = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
283 |
(t, [] |> Term.add_free_names t |> maps all_prefixes_of) |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
284 |
|> fold (fn ((s, i), T) => fn (t', taken) => |
57729 | 285 |
let val s' = singleton (Name.variant_list taken) s in |
286 |
((if fastype_of t' = HOLogic.boolT then HOLogic.all_const |
|
287 |
else Logic.all_const) T |
|
288 |
$ Abs (s', T, abstract_over (Var ((s, i), T), t')), |
|
289 |
s' :: taken) |
|
290 |
end) |
|
60924
610794dff23c
tuned signature, in accordance to sortBy in Scala;
wenzelm
parents:
60448
diff
changeset
|
291 |
(Term.add_vars t [] |> sort_by (fst o fst)) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
292 |
|> fst |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
293 |
|
62095 | 294 |
fun cartouche_term ctxt = close_form #> hackish_string_of_term ctxt #> cartouche |
295 |
fun cartouche_thm ctxt = cartouche_term ctxt o Thm.prop_of |
|
48394
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48332
diff
changeset
|
296 |
|
54084 | 297 |
(* TODO: rewrite to use nets and/or to reuse existing data structures *) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
298 |
fun clasimpset_rule_table_of ctxt = |
53502 | 299 |
let val simps = ctxt |> simpset_of |> dest_ss |> #simps in |
300 |
if length simps >= max_simps_for_clasimpset then |
|
301 |
Termtab.empty |
|
302 |
else |
|
303 |
let |
|
59582 | 304 |
fun add stature th = Termtab.update (normalize_vars (Thm.prop_of th), stature) |
57468 | 305 |
|
60943 | 306 |
val {safeIs, (* safeEs, *) unsafeIs, (* unsafeEs, *) ...} = |
307 |
ctxt |> claset_of |> Classical.rep_cs |
|
60945 | 308 |
val intros = map #1 (Item_Net.content safeIs @ Item_Net.content unsafeIs) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
309 |
(* Add once it is used: |
60943 | 310 |
val elims = Item_Net.content safeEs @ Item_Net.content unsafeEs |
53502 | 311 |
|> map Classical.classical_rule |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
312 |
*) |
57729 | 313 |
val specs = Spec_Rules.get ctxt |
314 |
val (rec_defs, nonrec_defs) = specs |
|
71179 | 315 |
|> filter (Spec_Rules.is_equational o #rough_classification) |
316 |
|> maps #rules |
|
59582 | 317 |
|> List.partition (is_rec_def o Thm.prop_of) |
57729 | 318 |
val spec_intros = specs |
71179 | 319 |
|> filter (Spec_Rules.is_relational o #rough_classification) |
320 |
|> maps #rules |
|
53502 | 321 |
in |
57468 | 322 |
Termtab.empty |
323 |
|> fold (add Simp o snd) simps |
|
324 |
|> fold (add Rec_Def) rec_defs |
|
325 |
|> fold (add Non_Rec_Def) nonrec_defs |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
326 |
(* Add once it is used: |
57468 | 327 |
|> fold (add Elim) elims |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
328 |
*) |
57468 | 329 |
|> fold (add Intro) intros |
330 |
|> fold (add Inductive) spec_intros |
|
53502 | 331 |
end |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
332 |
end |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
333 |
|
69597 | 334 |
fun normalize_eq (\<^const>\<open>Trueprop\<close> $ (t as (t0 as Const (\<^const_name>\<open>HOL.eq\<close>, _)) $ t1 $ t2)) = |
67560 | 335 |
if is_less_equal (Term_Ord.fast_term_ord (t1, t2)) then t else t0 $ t2 $ t1 |
69597 | 336 |
| normalize_eq (\<^const>\<open>Trueprop\<close> $ (t as \<^const>\<open>Not\<close> |
69593 | 337 |
$ ((t0 as Const (\<^const_name>\<open>HOL.eq\<close>, _)) $ t1 $ t2))) = |
67560 | 338 |
if is_less_equal (Term_Ord.fast_term_ord (t1, t2)) then t else HOLogic.mk_not (t0 $ t2 $ t1) |
69593 | 339 |
| normalize_eq (Const (\<^const_name>\<open>Pure.eq\<close>, Type (_, [T, _])) $ t1 $ t2) = |
67560 | 340 |
(if is_less_equal (Term_Ord.fast_term_ord (t1, t2)) then (t1, t2) else (t2, t1)) |
54092 | 341 |
|> (fn (t1, t2) => HOLogic.eq_const T $ t1 $ t2) |
50481
5d147d492792
push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents:
50442
diff
changeset
|
342 |
| normalize_eq t = t |
5d147d492792
push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents:
50442
diff
changeset
|
343 |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
344 |
fun if_thm_before th th' = |
65458 | 345 |
if Context.subthy_id (apply2 Thm.theory_id (th, th')) then th else th' |
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
346 |
|
57468 | 347 |
(* Hack: Conflate the facts about a class as seen from the outside with the corresponding low-level |
348 |
facts, so that MaSh can learn from the low-level proofs. *) |
|
50734 | 349 |
fun un_class_ify s = |
55286 | 350 |
(case first_field "_class" s of |
50734 | 351 |
SOME (pref, suf) => [s, pref ^ suf] |
55286 | 352 |
| NONE => [s]) |
50734 | 353 |
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
354 |
fun build_name_tables name_of facts = |
50733 | 355 |
let |
59582 | 356 |
fun cons_thm (_, th) = Termtab.cons_list (normalize_vars (normalize_eq (Thm.prop_of th)), th) |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
357 |
fun add_plain canon alias = |
54092 | 358 |
Symtab.update (Thm.get_name_hint alias, name_of (if_thm_before canon alias)) |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
359 |
fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases |
54092 | 360 |
fun add_inclass (name, target) = fold (fn s => Symtab.update (s, target)) (un_class_ify name) |
50815 | 361 |
val prop_tab = fold cons_thm facts Termtab.empty |
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
362 |
val plain_name_tab = Termtab.fold add_plains prop_tab Symtab.empty |
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
363 |
val inclass_name_tab = Symtab.fold add_inclass plain_name_tab Symtab.empty |
57108
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
56245
diff
changeset
|
364 |
in |
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
56245
diff
changeset
|
365 |
(plain_name_tab, inclass_name_tab) |
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
56245
diff
changeset
|
366 |
end |
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
367 |
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
368 |
fun fact_distinct eq facts = |
61322 | 369 |
fold (fn (i, fact as (_, th)) => |
370 |
Net.insert_term_safe (eq o apply2 (normalize_eq o Thm.prop_of o snd o snd)) |
|
371 |
(normalize_eq (Thm.prop_of th), (i, fact))) |
|
372 |
(tag_list 0 facts) Net.empty |
|
54076 | 373 |
|> Net.entries |
61322 | 374 |
|> sort (int_ord o apply2 fst) |
375 |
|> map snd |
|
53504 | 376 |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
377 |
fun struct_induct_rule_on th = |
59582 | 378 |
(case Logic.strip_horn (Thm.prop_of th) of |
69597 | 379 |
(prems, \<^const>\<open>Trueprop\<close> $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) => |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
380 |
if not (is_TVar ind_T) andalso length prems > 1 andalso |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
381 |
exists (exists_subterm (curry (op aconv) p)) prems andalso |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
382 |
not (exists (exists_subterm (curry (op aconv) a)) prems) then |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
383 |
SOME (p_name, ind_T) |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
384 |
else |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
385 |
NONE |
55286 | 386 |
| _ => NONE) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
387 |
|
50586 | 388 |
val instantiate_induct_timeout = seconds 0.01 |
389 |
||
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
390 |
fun instantiate_induct_rule ctxt concl_prop p_name ((name, stature), th) ind_x = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
391 |
let |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
392 |
fun varify_noninducts (t as Free (s, T)) = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
393 |
if (s, T) = ind_x orelse can dest_funT T then t else Var ((s, 0), T) |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
394 |
| varify_noninducts t = t |
57468 | 395 |
|
57729 | 396 |
val p_inst = concl_prop |
397 |
|> map_aterms varify_noninducts |
|
398 |
|> close_form |
|
399 |
|> lambda (Free ind_x) |
|
400 |
|> hackish_string_of_term ctxt |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
401 |
in |
57729 | 402 |
((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]", stature), |
59755 | 403 |
th |> Rule_Insts.read_instantiate ctxt [(((p_name, 0), Position.none), p_inst)] []) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
404 |
end |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
405 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
406 |
fun type_match thy (T1, T2) = |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
407 |
(Sign.typ_match thy (T2, T1) Vartab.empty; true) |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
408 |
handle Type.TYPE_MATCH => false |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
409 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
410 |
fun instantiate_if_induct_rule ctxt stmt stmt_xs (ax as (_, th)) = |
55286 | 411 |
(case struct_induct_rule_on th of |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
412 |
SOME (p_name, ind_T) => |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
413 |
let val thy = Proof_Context.theory_of ctxt in |
55286 | 414 |
stmt_xs |
415 |
|> filter (fn (_, T) => type_match thy (T, ind_T)) |
|
62519 | 416 |
|> map_filter (try (Timeout.apply instantiate_induct_timeout |
57729 | 417 |
(instantiate_induct_rule ctxt stmt p_name ax))) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
418 |
end |
55286 | 419 |
| NONE => [ax]) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
420 |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
421 |
fun external_frees t = |
55948 | 422 |
[] |> Term.add_frees t |> filter_out (Name.is_internal o fst) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
423 |
|
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
424 |
fun instantiate_inducts ctxt hyp_ts concl_t = |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
425 |
let |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
426 |
val ind_stmt = |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
427 |
(hyp_ts |> filter_out (null o external_frees), concl_t) |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
428 |
|> Logic.list_implies |> Object_Logic.atomize_term ctxt |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
429 |
val ind_stmt_xs = external_frees ind_stmt |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
430 |
in |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
431 |
maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs) |
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
432 |
end |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
433 |
|
73979 | 434 |
fun fact_of_lazy_fact ((name, stature), th) = ((name (), stature), th) |
51004
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents:
50815
diff
changeset
|
435 |
|
53532
4ad9599a0847
disable some checks for huge background theories
blanchet
parents:
53531
diff
changeset
|
436 |
fun fact_count facts = Facts.fold_static (K (Integer.add 1)) facts 0 |
4ad9599a0847
disable some checks for huge background theories
blanchet
parents:
53531
diff
changeset
|
437 |
|
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
438 |
fun is_useful_unnamed_local_fact ctxt = |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
439 |
let |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
440 |
val thy = Proof_Context.theory_of ctxt |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
441 |
val global_facts = Global_Theory.facts_of thy |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
442 |
val local_facts = Proof_Context.facts_of ctxt |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
443 |
val named_locals = Facts.dest_static true [global_facts] local_facts |
66613 | 444 |
|> maps (map (normalize_eq o Thm.prop_of) o snd) |
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
445 |
in |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
446 |
fn th => |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
447 |
not (Thm.has_name_hint th) andalso |
66613 | 448 |
not (member (op aconv) named_locals (normalize_eq (Thm.prop_of th))) |
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
449 |
end |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
450 |
|
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
451 |
fun all_facts ctxt generous keywords add_ths chained css = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
452 |
let |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
453 |
val thy = Proof_Context.theory_of ctxt |
64413 | 454 |
val transfer = Global_Theory.transfer_theories thy |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
455 |
val global_facts = Global_Theory.facts_of thy |
53532
4ad9599a0847
disable some checks for huge background theories
blanchet
parents:
53531
diff
changeset
|
456 |
val is_too_complex = |
57729 | 457 |
if generous orelse fact_count global_facts >= max_facts_for_complex_check then K false |
458 |
else is_too_complex |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
459 |
val local_facts = Proof_Context.facts_of ctxt |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
460 |
val assms = Assumption.all_assms_of ctxt |
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
461 |
val named_locals = Facts.dest_static true [global_facts] local_facts |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
462 |
val unnamed_locals = |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
463 |
Facts.props local_facts |
63337 | 464 |
|> map #1 |
58089
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
465 |
|> filter (is_useful_unnamed_local_fact ctxt) |
20e76da3a0ef
gracefully reconstruct Isar proofs in scenarios such as 'using f unfolding g', where backticks can't be used to refer to the unfolded version of 'f' (for some reason)
blanchet
parents:
58011
diff
changeset
|
466 |
|> map (pair "" o single) |
57468 | 467 |
|
57729 | 468 |
val full_space = Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts) |
57468 | 469 |
|
56141
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents:
56140
diff
changeset
|
470 |
fun add_facts global foldx facts = |
54081 | 471 |
foldx (fn (name0, ths) => fn accum => |
50512
c283bc0a8f1a
tweaked which facts are included for MaSh evaluations
blanchet
parents:
50511
diff
changeset
|
472 |
if name0 <> "" andalso |
59916 | 473 |
(Long_Name.is_hidden (Facts.intern facts name0) orelse |
58226
04faf6dc262e
never include hidden names -- these cannot be referenced afterward
blanchet
parents:
58089
diff
changeset
|
474 |
((Facts.is_concealed facts name0 orelse |
04faf6dc262e
never include hidden names -- these cannot be referenced afterward
blanchet
parents:
58089
diff
changeset
|
475 |
(not generous andalso is_blacklisted_or_something name0)) andalso |
04faf6dc262e
never include hidden names -- these cannot be referenced afterward
blanchet
parents:
58089
diff
changeset
|
476 |
forall (not o member Thm.eq_thm_prop add_ths) ths)) then |
54081 | 477 |
accum |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
478 |
else |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
479 |
let |
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
480 |
val n = length ths |
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
481 |
val collection = n > 1 |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
482 |
val dotted_name = length (Long_Name.explode name0) > 2 (* ignore theory name *) |
57468 | 483 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
484 |
fun check_thms a = |
55286 | 485 |
(case try (Proof_Context.get_thms ctxt) a of |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
486 |
NONE => false |
55286 | 487 |
| SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
488 |
in |
61054 | 489 |
snd (fold_rev (fn th0 => fn (j, accum) => |
490 |
let val th = transfer th0 in |
|
491 |
(j - 1, |
|
492 |
if not (member Thm.eq_thm_prop add_ths th) andalso |
|
493 |
(is_likely_tautology_too_meta_or_too_technical th orelse |
|
494 |
is_too_complex (Thm.prop_of th)) then |
|
495 |
accum |
|
496 |
else |
|
497 |
let |
|
498 |
fun get_name () = |
|
499 |
if name0 = "" orelse name0 = local_thisN then |
|
62095 | 500 |
cartouche_thm ctxt th |
61054 | 501 |
else |
502 |
let val short_name = Facts.extern ctxt facts name0 in |
|
503 |
if check_thms short_name then |
|
504 |
short_name |
|
505 |
else |
|
506 |
let val long_name = Name_Space.extern ctxt full_space name0 in |
|
507 |
if check_thms long_name then |
|
508 |
long_name |
|
509 |
else |
|
510 |
name0 |
|
511 |
end |
|
512 |
end |
|
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
513 |
|> make_name keywords collection j |
61054 | 514 |
val stature = stature_of_thm global assms chained css name0 th |
515 |
val new = ((get_name, stature), th) |
|
516 |
in |
|
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
517 |
(if collection then apsnd o apsnd |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
518 |
else if dotted_name then apsnd o apfst |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
519 |
else apfst) (cons new) accum |
61054 | 520 |
end) |
521 |
end) ths (n, accum)) |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
522 |
end) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
523 |
in |
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
524 |
(* Names like "xxx" are preferred to "xxx.yyy", which are preferred to "xxx(666)" and the like. |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
525 |
"Preferred" means put to the front of the list. *) |
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
526 |
([], ([], [])) |
57729 | 527 |
|> add_facts false fold local_facts (unnamed_locals @ named_locals) |
528 |
|> add_facts true Facts.fold_static global_facts global_facts |
|
73980
291f7b5fc7c9
prefer 'xxx' to 'xxx.yyy' to 'xxx(2)' in Sledgehammer, to some extent
blanchet
parents:
73979
diff
changeset
|
529 |
||> op @ |> op @ |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
530 |
end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
531 |
|
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
532 |
fun nearly_all_facts ctxt inst_inducts {add, del, only} keywords css chained hyp_ts |
73939
9231ea46e041
promoted "sledgehammer_instantiate_inducts" to proper option "induction_rules"
desharna
parents:
72342
diff
changeset
|
533 |
concl_t = |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
534 |
if only andalso null add then |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
535 |
[] |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
536 |
else |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
537 |
let |
57729 | 538 |
val chained = chained |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th]) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
539 |
in |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
540 |
(if only then |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
541 |
maps (map (fn ((name, stature), th) => ((K name, stature), th)) |
58919 | 542 |
o fact_of_ref ctxt keywords chained css) add |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
543 |
else |
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
544 |
let |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58928
diff
changeset
|
545 |
val (add, del) = apply2 (Attrib.eval_thms ctxt) (add, del) |
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
546 |
val facts = |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
547 |
all_facts ctxt false keywords add chained css |
54143
18def1c73c79
make sure add: doesn't add duplicates, and works for [no_atp] facts
blanchet
parents:
54142
diff
changeset
|
548 |
|> filter_out ((member Thm.eq_thm_prop del orf |
69593 | 549 |
(Named_Theorems.member ctxt \<^named_theorems>\<open>no_atp\<close> andf |
57963 | 550 |
not o member Thm.eq_thm_prop add)) o snd) |
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
551 |
in |
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
552 |
facts |
48292 | 553 |
end) |
73940
f58108b7a60c
refactored Sledgehammer option "induction_rules"
desharna
parents:
73939
diff
changeset
|
554 |
|> inst_inducts ? instantiate_inducts ctxt hyp_ts concl_t |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
555 |
end |
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
556 |
|
54142
5f3c1b445253
no fact subsumption -- this only confuses later code, e.g. 'add:'
blanchet
parents:
54112
diff
changeset
|
557 |
fun drop_duplicate_facts facts = |
54112
9c0f464d1854
drop only real duplicates, not subsumed facts -- this confuses MaSh
blanchet
parents:
54092
diff
changeset
|
558 |
let val num_facts = length facts in |
54142
5f3c1b445253
no fact subsumption -- this only confuses later code, e.g. 'add:'
blanchet
parents:
54112
diff
changeset
|
559 |
facts |> num_facts <= max_facts_for_duplicates ? fact_distinct (op aconv) |
54112
9c0f464d1854
drop only real duplicates, not subsumed facts -- this confuses MaSh
blanchet
parents:
54092
diff
changeset
|
560 |
end |
9c0f464d1854
drop only real duplicates, not subsumed facts -- this confuses MaSh
blanchet
parents:
54092
diff
changeset
|
561 |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
562 |
end; |