| author | blanchet |
| Wed, 06 Jan 2016 13:04:31 +0100 | |
| changeset 62081 | fd18b51bdc55 |
| parent 61877 | 276ad4354069 |
| child 62093 | bd73a2279fcd |
| 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 |
|
|
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
|
13 |
type raw_fact = ((unit -> string) * stature) * thm |
|
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 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
21 |
val instantiate_inducts : bool Config.T |
| 48292 | 22 |
val no_fact_override : fact_override |
| 55212 | 23 |
|
| 58919 | 24 |
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
|
25 |
Facts.ref * Token.src list -> ((string * stature) * thm) list |
|
50047
45684acf0b94
thread context correctly when printing backquoted facts
blanchet
parents:
49981
diff
changeset
|
26 |
val backquote_thm : Proof.context -> thm -> string |
|
50511
8825c36cb1ce
don't query blacklisted theorems in evaluation driver
blanchet
parents:
50510
diff
changeset
|
27 |
val is_blacklisted_or_something : Proof.context -> bool -> 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
|
28 |
val clasimpset_rule_table_of : Proof.context -> status Termtab.table |
| 55212 | 29 |
val build_name_tables : (thm -> string) -> ('a * thm) list ->
|
30 |
string Symtab.table * string Symtab.table |
|
| 61322 | 31 |
val fact_distinct : (term * term -> bool) -> ('a * thm) list -> ('a * thm) list
|
| 55212 | 32 |
val maybe_instantiate_inducts : Proof.context -> term list -> term -> |
33 |
(((unit -> string) * 'a) * thm) list -> (((unit -> string) * 'a) * thm) list |
|
|
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
|
34 |
val fact_of_raw_fact : raw_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
|
35 |
val is_useful_unnamed_local_fact : Proof.context -> thm -> bool |
| 55212 | 36 |
|
| 58919 | 37 |
val all_facts : Proof.context -> bool -> bool -> Keyword.keywords -> thm list -> thm list -> |
| 55212 | 38 |
status Termtab.table -> raw_fact list |
| 58919 | 39 |
val nearly_all_facts : Proof.context -> bool -> fact_override -> Keyword.keywords -> |
| 55212 | 40 |
status Termtab.table -> thm list -> term list -> term -> raw_fact list |
|
54142
5f3c1b445253
no fact subsumption -- this only confuses later code, e.g. 'add:'
blanchet
parents:
54112
diff
changeset
|
41 |
val drop_duplicate_facts : raw_fact list -> raw_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
|
42 |
end; |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
43 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
44 |
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
|
45 |
struct |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
46 |
|
|
50495
bd9a0028b063
better tautology check -- don't reject "prod_cases3" for example
blanchet
parents:
50490
diff
changeset
|
47 |
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
|
48 |
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
|
49 |
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
|
50 |
|
|
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 raw_fact = ((unit -> string) * stature) * thm |
|
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
|
52 |
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
|
53 |
|
| 48292 | 54 |
type fact_override = |
|
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
55 |
{add : (Facts.ref * Token.src list) list,
|
|
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57983
diff
changeset
|
56 |
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
|
57 |
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
|
58 |
|
|
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
|
59 |
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
|
60 |
|
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
61 |
(* gracefully handle huge background theories *) |
|
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
62 |
val max_facts_for_duplicates = 50000 |
|
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
63 |
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
|
64 |
val max_simps_for_clasimpset = 10000 |
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
65 |
|
|
53511
3ea6b9461c55
got rid of another slowdown factor in relevance filter
blanchet
parents:
53510
diff
changeset
|
66 |
(* experimental feature *) |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
67 |
val instantiate_inducts = |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
68 |
Attrib.setup_config_bool @{binding sledgehammer_instantiate_inducts} (K false)
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
69 |
|
| 48292 | 70 |
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
|
71 |
|
| 58919 | 72 |
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
|
73 |
Keyword.is_literal keywords s orelse |
| 50239 | 74 |
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
|
75 |
|
| 58919 | 76 |
fun make_name keywords multi j name = |
77 |
(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
|
78 |
(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
|
79 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
80 |
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
|
81 |
| 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
|
82 |
| 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
|
83 |
|
| 61646 | 84 |
val backquote = enclose "\<open>" "\<close>" |
|
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 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
86 |
(* unfolding these can yield really huge terms *) |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
87 |
val risky_defs = @{thms Bit0_def Bit1_def}
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
88 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
89 |
fun is_rec_eq lhs = Term.exists_subterm (curry (op =) (head_of lhs)) |
| 57468 | 90 |
|
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
91 |
fun is_rec_def (@{const Trueprop} $ t) = is_rec_def t
|
| 56245 | 92 |
| is_rec_def (@{const Pure.imp} $ _ $ t2) = is_rec_def t2
|
93 |
| is_rec_def (Const (@{const_name Pure.eq}, _) $ 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
|
94 |
| is_rec_def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) = is_rec_eq 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
|
95 |
| 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
|
96 |
|
| 59582 | 97 |
fun is_assum assms th = exists (fn ct => Thm.prop_of th aconv Thm.term_of ct) assms |
| 48396 | 98 |
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
|
99 |
|
| 48396 | 100 |
fun scope_of_thm global assms chained th = |
101 |
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
|
102 |
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
|
103 |
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
|
104 |
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
|
105 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
106 |
val may_be_induction = |
| 57729 | 107 |
exists_subterm (fn Var (_, Type (@{type_name fun}, [_, T])) => body_type T = @{typ bool}
|
108 |
| _ => 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
|
109 |
|
| 54077 | 110 |
(* TODO: get rid of *) |
| 53501 | 111 |
fun normalize_vars t = |
112 |
let |
|
113 |
fun normT (Type (s, Ts)) = fold_map normT Ts #>> curry Type s |
|
114 |
| normT (TVar (z as (_, S))) = |
|
115 |
(fn ((knownT, nT), accum) => |
|
| 55286 | 116 |
(case find_index (equal z) knownT of |
| 53501 | 117 |
~1 => (TVar ((Name.uu, nT), S), ((z :: knownT, nT + 1), accum)) |
| 55286 | 118 |
| j => (TVar ((Name.uu, nT - j - 1), S), ((knownT, nT), accum)))) |
| 53501 | 119 |
| normT (T as TFree _) = pair T |
| 57468 | 120 |
|
| 53501 | 121 |
fun norm (t $ u) = norm t ##>> norm u #>> op $ |
122 |
| norm (Const (s, T)) = normT T #>> curry Const s |
|
| 57729 | 123 |
| norm (Var (z as (_, T))) = normT T |
| 53501 | 124 |
#> (fn (T, (accumT, (known, n))) => |
| 57729 | 125 |
(case find_index (equal z) known of |
126 |
~1 => (Var ((Name.uu, n), T), (accumT, (z :: known, n + 1))) |
|
127 |
| j => (Var ((Name.uu, n - j - 1), T), (accumT, (known, n))))) |
|
128 |
| norm (Abs (_, T, t)) = norm t ##>> normT T #>> (fn (t, T) => Abs (Name.uu, T, t)) |
|
| 53501 | 129 |
| norm (Bound j) = pair (Bound j) |
130 |
| norm (Free (s, T)) = normT T #>> curry Free s |
|
131 |
in fst (norm t (([], 0), ([], 0))) end |
|
132 |
||
| 48396 | 133 |
fun status_of_thm css name th = |
| 54081 | 134 |
if Termtab.is_empty css then |
135 |
General |
|
136 |
else |
|
| 59582 | 137 |
let val t = Thm.prop_of th in |
| 54081 | 138 |
(* FIXME: use structured name *) |
139 |
if String.isSubstring ".induct" name andalso may_be_induction t then |
|
140 |
Induction |
|
141 |
else |
|
142 |
let val t = normalize_vars t in |
|
| 55286 | 143 |
(case Termtab.lookup css t of |
| 54081 | 144 |
SOME status => status |
145 |
| NONE => |
|
146 |
let val concl = Logic.strip_imp_concl t in |
|
| 55286 | 147 |
(case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of |
| 54081 | 148 |
SOME lrhss => |
149 |
let |
|
150 |
val prems = Logic.strip_imp_prems t |
|
151 |
val t' = Logic.list_implies (prems, Logic.mk_equals lrhss) |
|
152 |
in |
|
153 |
Termtab.lookup css t' |> the_default General |
|
154 |
end |
|
| 55286 | 155 |
| NONE => General) |
156 |
end) |
|
| 54081 | 157 |
end |
158 |
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
|
159 |
|
| 48396 | 160 |
fun stature_of_thm global assms chained css name th = |
161 |
(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
|
162 |
|
| 58919 | 163 |
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
|
164 |
let |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
165 |
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
|
166 |
val bracket = |
|
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
wenzelm
parents:
61646
diff
changeset
|
167 |
implode (map (enclose "[" "]" o Pretty.unformatted_string_of o Token.pretty_src ctxt) args) |
| 57468 | 168 |
|
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
169 |
fun nth_name j = |
| 55286 | 170 |
(case xref of |
| 59433 | 171 |
Facts.Fact s => backquote (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
|
172 |
| Facts.Named (("", _), _) => "[" ^ bracket ^ "]"
|
| 58919 | 173 |
| 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
|
174 |
| Facts.Named ((name, _), SOME intervals) => |
| 58919 | 175 |
make_name keywords true |
| 55286 | 176 |
(nth (maps (explode_interval (length ths)) intervals) j) name ^ bracket) |
| 57468 | 177 |
|
| 48396 | 178 |
fun add_nth th (j, rest) = |
179 |
let val name = nth_name j in |
|
| 57729 | 180 |
(j + 1, ((name, stature_of_thm false [] chained css name th), th) :: rest) |
| 48396 | 181 |
end |
| 57468 | 182 |
in |
183 |
(0, []) |> fold add_nth ths |> snd |
|
184 |
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
|
185 |
|
| 57729 | 186 |
(* Reject theorems with names like "List.filter.filter_list_def" or "Accessible_Part.acc.defs", as |
187 |
these are definitions arising from packages. *) |
|
| 50736 | 188 |
fun is_package_def s = |
189 |
let val ss = Long_Name.explode s in |
|
190 |
length ss > 2 andalso not (hd ss = "local") andalso |
|
191 |
exists (fn suf => String.isSuffix suf s) |
|
| 57729 | 192 |
["_case_def", "_rec_def", "_size_def", "_size_overloaded_def"] |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
193 |
end |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
194 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
195 |
(* FIXME: put other record thms here, or declare as "no_atp" *) |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
196 |
fun multi_base_blacklist ctxt ho_atp = |
| 57729 | 197 |
["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
|
198 |
"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
|
199 |
"nibble.distinct"] |
| 57729 | 200 |
|> not (ho_atp orelse Config.get ctxt instantiate_inducts) ? append ["induct", "inducts"] |
|
48440
159e320ef851
identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents:
48438
diff
changeset
|
201 |
|> 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
|
202 |
|
| 57729 | 203 |
(* 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
|
204 |
val max_apply_depth = 18 |
|
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
205 |
|
|
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
206 |
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
|
207 |
| 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
|
208 |
| apply_depth _ = 0 |
|
1eb7c65b526c
reverted f99ee3adb81d -- that old logic seems to make a difference still today
blanchet
parents:
53513
diff
changeset
|
209 |
|
| 53533 | 210 |
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
|
211 |
|
|
48667
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
212 |
(* FIXME: Ad hoc list *) |
|
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
213 |
val technical_prefixes = |
| 57729 | 214 |
["ATP", "Code_Evaluation", "Datatype", "Enum", "Lazy_Sequence", "Limited_Sequence", "Meson", |
215 |
"Metis", "Nitpick", "Quickcheck_Random", "Quickcheck_Exhaustive", "Quickcheck_Narrowing", |
|
| 50736 | 216 |
"Random_Sequence", "Sledgehammer", "SMT"] |
|
48667
ac58317ef11f
rule out same "technical" theories for MePo as for MaSh
blanchet
parents:
48530
diff
changeset
|
217 |
|> 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
|
218 |
|
| 57729 | 219 |
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
|
220 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
221 |
(* FIXME: make more reliable *) |
| 53507 | 222 |
val sep_class_sep = Long_Name.separator ^ "class" ^ Long_Name.separator |
| 57468 | 223 |
|
| 57729 | 224 |
fun is_low_level_class_const s = |
| 53507 | 225 |
s = @{const_name equal_class.equal} 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
|
226 |
|
| 60448 | 227 |
val sep_that = Long_Name.separator ^ Auto_Bind.thatN |
| 54081 | 228 |
val skolem_thesis = Name.skolem Auto_Bind.thesisN |
229 |
||
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
230 |
fun is_that_fact th = |
| 59582 | 231 |
exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (Thm.prop_of th) |
| 54081 | 232 |
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
|
233 |
|
| 53513 | 234 |
datatype interest = Deal_Breaker | Interesting | Boring |
235 |
||
236 |
fun combine_interests Deal_Breaker _ = Deal_Breaker |
|
237 |
| combine_interests _ Deal_Breaker = Deal_Breaker |
|
238 |
| combine_interests Interesting _ = Interesting |
|
239 |
| combine_interests _ Interesting = Interesting |
|
240 |
| combine_interests Boring Boring = Boring |
|
241 |
||
|
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
|
242 |
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
|
243 |
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
|
244 |
|
|
50523
0799339fea0f
get rid of some junk facts in the MaSh evaluation driver
blanchet
parents:
50512
diff
changeset
|
245 |
fun is_likely_tautology_too_meta_or_too_technical th = |
| 48406 | 246 |
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
|
247 |
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
|
248 |
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
|
249 |
| 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
|
250 |
| is_interesting_subterm _ = false |
| 57468 | 251 |
|
| 53513 | 252 |
fun interest_of_bool t = |
| 57729 | 253 |
if exists_Const ((is_technical_const o fst) orf (is_low_level_class_const o fst) orf |
254 |
type_has_top_sort o snd) t then |
|
| 53513 | 255 |
Deal_Breaker |
256 |
else if exists_type (exists_subtype (curry (op =) @{typ prop})) t orelse
|
|
| 57729 | 257 |
not (exists_subterm is_interesting_subterm t) then |
| 53513 | 258 |
Boring |
259 |
else |
|
260 |
Interesting |
|
| 57468 | 261 |
|
| 53513 | 262 |
fun interest_of_prop _ (@{const Trueprop} $ t) = interest_of_bool t
|
| 56245 | 263 |
| interest_of_prop Ts (@{const Pure.imp} $ t $ u) =
|
| 53513 | 264 |
combine_interests (interest_of_prop Ts t) (interest_of_prop Ts u) |
| 56245 | 265 |
| interest_of_prop Ts (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) =
|
| 54040 | 266 |
if type_has_top_sort T then Deal_Breaker else interest_of_prop (T :: Ts) t |
| 56245 | 267 |
| interest_of_prop Ts ((t as Const (@{const_name Pure.all}, _)) $ u) =
|
| 53513 | 268 |
interest_of_prop Ts (t $ eta_expand Ts u 1) |
| 56245 | 269 |
| interest_of_prop _ (Const (@{const_name Pure.eq}, _) $ t $ u) =
|
| 53513 | 270 |
combine_interests (interest_of_bool t) (interest_of_bool u) |
271 |
| interest_of_prop _ _ = Deal_Breaker |
|
| 57468 | 272 |
|
| 59582 | 273 |
val t = Thm.prop_of th |
| 48406 | 274 |
in |
| 57729 | 275 |
(interest_of_prop [] t <> Interesting andalso not (Thm.eq_thm_prop (@{thm ext}, th))) orelse
|
| 53507 | 276 |
is_that_fact th |
|
48438
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48437
diff
changeset
|
277 |
end |
|
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents:
48437
diff
changeset
|
278 |
|
| 53512 | 279 |
fun is_blacklisted_or_something ctxt ho_atp = |
| 57468 | 280 |
let val blist = multi_base_blacklist ctxt ho_atp in |
281 |
fn name => is_package_def name orelse exists (fn s => String.isSuffix s name) blist |
|
282 |
end |
|
| 50510 | 283 |
|
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
284 |
(* 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
|
285 |
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
|
286 |
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
|
287 |
prefixes of all free variables. In the worse case scenario, where the fact |
| 49981 | 288 |
won't be resolved correctly, the user can fix it manually, e.g., by giving a |
289 |
name to the offending fact. *) |
|
| 57729 | 290 |
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
|
291 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
292 |
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
|
293 |
(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
|
294 |
|> fold (fn ((s, i), T) => fn (t', taken) => |
| 57729 | 295 |
let val s' = singleton (Name.variant_list taken) s in |
296 |
((if fastype_of t' = HOLogic.boolT then HOLogic.all_const |
|
297 |
else Logic.all_const) T |
|
298 |
$ Abs (s', T, abstract_over (Var ((s, i), T), t')), |
|
299 |
s' :: taken) |
|
300 |
end) |
|
|
60924
610794dff23c
tuned signature, in accordance to sortBy in Scala;
wenzelm
parents:
60448
diff
changeset
|
301 |
(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
|
302 |
|> fst |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
303 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51160
diff
changeset
|
304 |
fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote |
| 59582 | 305 |
fun backquote_thm ctxt = backquote_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
|
306 |
|
| 54084 | 307 |
(* 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
|
308 |
fun clasimpset_rule_table_of ctxt = |
| 53502 | 309 |
let val simps = ctxt |> simpset_of |> dest_ss |> #simps in |
310 |
if length simps >= max_simps_for_clasimpset then |
|
311 |
Termtab.empty |
|
312 |
else |
|
313 |
let |
|
| 59582 | 314 |
fun add stature th = Termtab.update (normalize_vars (Thm.prop_of th), stature) |
| 57468 | 315 |
|
| 60943 | 316 |
val {safeIs, (* safeEs, *) unsafeIs, (* unsafeEs, *) ...} =
|
317 |
ctxt |> claset_of |> Classical.rep_cs |
|
| 60945 | 318 |
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
|
319 |
(* Add once it is used: |
| 60943 | 320 |
val elims = Item_Net.content safeEs @ Item_Net.content unsafeEs |
| 53502 | 321 |
|> 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
|
322 |
*) |
| 57729 | 323 |
val specs = Spec_Rules.get ctxt |
324 |
val (rec_defs, nonrec_defs) = specs |
|
325 |
|> filter (curry (op =) Spec_Rules.Equational o fst) |
|
326 |
|> maps (snd o snd) |
|
327 |
|> filter_out (member Thm.eq_thm_prop risky_defs) |
|
| 59582 | 328 |
|> List.partition (is_rec_def o Thm.prop_of) |
| 57729 | 329 |
val spec_intros = specs |
330 |
|> filter (member (op =) [Spec_Rules.Inductive, Spec_Rules.Co_Inductive] o fst) |
|
331 |
|> maps (snd o snd) |
|
| 53502 | 332 |
in |
| 57468 | 333 |
Termtab.empty |
334 |
|> fold (add Simp o snd) simps |
|
335 |
|> fold (add Rec_Def) rec_defs |
|
336 |
|> 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
|
337 |
(* Add once it is used: |
| 57468 | 338 |
|> 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
|
339 |
*) |
| 57468 | 340 |
|> fold (add Intro) intros |
341 |
|> fold (add Inductive) spec_intros |
|
| 53502 | 342 |
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
|
343 |
end |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
344 |
|
| 54092 | 345 |
fun normalize_eq (@{const Trueprop} $ (t as (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
|
346 |
if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else t0 $ t2 $ t1 |
|
347 |
| normalize_eq (@{const Trueprop} $ (t as @{const Not}
|
|
| 57729 | 348 |
$ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
|
| 54092 | 349 |
if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else HOLogic.mk_not (t0 $ t2 $ t1) |
| 56245 | 350 |
| normalize_eq (Const (@{const_name Pure.eq}, Type (_, [T, _])) $ t1 $ t2) =
|
| 54092 | 351 |
(if Term_Ord.fast_term_ord (t1, t2) <> GREATER then (t1, t2) else (t2, t1)) |
352 |
|> (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
|
353 |
| 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
|
354 |
|
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
355 |
fun if_thm_before th th' = |
| 60948 | 356 |
if Context.subthy_id (apply2 Thm.theory_id_of_thm (th, th')) then th else th' |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
357 |
|
| 57468 | 358 |
(* Hack: Conflate the facts about a class as seen from the outside with the corresponding low-level |
359 |
facts, so that MaSh can learn from the low-level proofs. *) |
|
| 50734 | 360 |
fun un_class_ify s = |
| 55286 | 361 |
(case first_field "_class" s of |
| 50734 | 362 |
SOME (pref, suf) => [s, pref ^ suf] |
| 55286 | 363 |
| NONE => [s]) |
| 50734 | 364 |
|
|
50735
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
blanchet
parents:
50734
diff
changeset
|
365 |
fun build_name_tables name_of facts = |
| 50733 | 366 |
let |
| 59582 | 367 |
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
|
368 |
fun add_plain canon alias = |
| 54092 | 369 |
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
|
370 |
fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases |
| 54092 | 371 |
fun add_inclass (name, target) = fold (fn s => Symtab.update (s, target)) (un_class_ify name) |
| 50815 | 372 |
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
|
373 |
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
|
374 |
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
|
375 |
in |
|
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
56245
diff
changeset
|
376 |
(plain_name_tab, inclass_name_tab) |
|
dc0b4f50e288
more generous max number of suggestions, for more safety
blanchet
parents:
56245
diff
changeset
|
377 |
end |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
378 |
|
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
379 |
fun fact_distinct eq facts = |
| 61322 | 380 |
fold (fn (i, fact as (_, th)) => |
381 |
Net.insert_term_safe (eq o apply2 (normalize_eq o Thm.prop_of o snd o snd)) |
|
382 |
(normalize_eq (Thm.prop_of th), (i, fact))) |
|
383 |
(tag_list 0 facts) Net.empty |
|
| 54076 | 384 |
|> Net.entries |
| 61322 | 385 |
|> sort (int_ord o apply2 fst) |
386 |
|> map snd |
|
| 53504 | 387 |
|
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
388 |
fun struct_induct_rule_on th = |
| 59582 | 389 |
(case Logic.strip_horn (Thm.prop_of th) of |
| 57729 | 390 |
(prems, @{const Trueprop} $ ((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
|
391 |
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
|
392 |
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
|
393 |
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
|
394 |
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
|
395 |
else |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
396 |
NONE |
| 55286 | 397 |
| _ => 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
|
398 |
|
| 50586 | 399 |
val instantiate_induct_timeout = seconds 0.01 |
400 |
||
|
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 |
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
|
402 |
let |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
403 |
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
|
404 |
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
|
405 |
| varify_noninducts t = t |
| 57468 | 406 |
|
| 57729 | 407 |
val p_inst = concl_prop |
408 |
|> map_aterms varify_noninducts |
|
409 |
|> close_form |
|
410 |
|> lambda (Free ind_x) |
|
411 |
|> 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
|
412 |
in |
| 57729 | 413 |
((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]", stature), |
| 59755 | 414 |
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
|
415 |
end |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
416 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
417 |
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
|
418 |
(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
|
419 |
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
|
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 instantiate_if_induct_rule ctxt stmt stmt_xs (ax as (_, th)) = |
| 55286 | 422 |
(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
|
423 |
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
|
424 |
let val thy = Proof_Context.theory_of ctxt in |
| 55286 | 425 |
stmt_xs |
426 |
|> filter (fn (_, T) => type_match thy (T, ind_T)) |
|
427 |
|> map_filter (try (TimeLimit.timeLimit instantiate_induct_timeout |
|
| 57729 | 428 |
(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
|
429 |
end |
| 55286 | 430 |
| 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
|
431 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
432 |
fun external_frees t = |
| 55948 | 433 |
[] |> 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
|
434 |
|
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
435 |
fun maybe_instantiate_inducts ctxt hyp_ts concl_t = |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
436 |
if Config.get ctxt instantiate_inducts then |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
437 |
let |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
438 |
val ind_stmt = |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
439 |
(hyp_ts |> filter_out (null o external_frees), concl_t) |
| 59970 | 440 |
|> Logic.list_implies |> Object_Logic.atomize_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
|
441 |
val ind_stmt_xs = external_frees ind_stmt |
| 57729 | 442 |
in |
443 |
maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs) |
|
444 |
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
|
445 |
else |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
446 |
I |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
447 |
|
|
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
|
448 |
fun fact_of_raw_fact ((name, stature), th) = ((name (), stature), th) |
|
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
|
449 |
|
|
53532
4ad9599a0847
disable some checks for huge background theories
blanchet
parents:
53531
diff
changeset
|
450 |
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
|
451 |
|
|
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
|
452 |
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
|
453 |
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
|
454 |
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
|
455 |
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
|
456 |
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
|
457 |
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
|
458 |
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
|
459 |
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
|
460 |
not (Thm.has_name_hint th) andalso |
|
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 |
forall (fn (_, ths) => not (member Thm.eq_thm_prop ths th)) named_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
|
462 |
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
|
463 |
|
| 58919 | 464 |
fun all_facts ctxt generous ho_atp keywords add_ths chained css = |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
465 |
let |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
466 |
val thy = Proof_Context.theory_of ctxt |
| 61054 | 467 |
val transfer = Global_Theory.transfer_theories thy; |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
468 |
val global_facts = Global_Theory.facts_of thy |
|
53532
4ad9599a0847
disable some checks for huge background theories
blanchet
parents:
53531
diff
changeset
|
469 |
val is_too_complex = |
| 57729 | 470 |
if generous orelse fact_count global_facts >= max_facts_for_complex_check then K false |
471 |
else is_too_complex |
|
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
472 |
val local_facts = Proof_Context.facts_of ctxt |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
473 |
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
|
474 |
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
|
475 |
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
|
476 |
Facts.props 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
|
477 |
|> 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
|
478 |
|> map (pair "" o single) |
| 57468 | 479 |
|
| 57729 | 480 |
val full_space = Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts) |
| 53512 | 481 |
val is_blacklisted_or_something = is_blacklisted_or_something ctxt ho_atp |
| 57468 | 482 |
|
|
56141
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents:
56140
diff
changeset
|
483 |
fun add_facts global foldx facts = |
| 54081 | 484 |
foldx (fn (name0, ths) => fn accum => |
|
50512
c283bc0a8f1a
tweaked which facts are included for MaSh evaluations
blanchet
parents:
50511
diff
changeset
|
485 |
if name0 <> "" andalso |
| 59916 | 486 |
(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
|
487 |
((Facts.is_concealed facts name0 orelse |
|
04faf6dc262e
never include hidden names -- these cannot be referenced afterward
blanchet
parents:
58089
diff
changeset
|
488 |
(not generous andalso is_blacklisted_or_something name0)) andalso |
|
04faf6dc262e
never include hidden names -- these cannot be referenced afterward
blanchet
parents:
58089
diff
changeset
|
489 |
forall (not o member Thm.eq_thm_prop add_ths) ths)) then |
| 54081 | 490 |
accum |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
491 |
else |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
492 |
let |
|
50485
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
493 |
val n = length ths |
|
3c6ac2da2f45
merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents:
50481
diff
changeset
|
494 |
val multi = n > 1 |
| 57468 | 495 |
|
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
496 |
fun check_thms a = |
| 55286 | 497 |
(case try (Proof_Context.get_thms ctxt) a of |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
498 |
NONE => false |
| 55286 | 499 |
| 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
|
500 |
in |
| 61054 | 501 |
snd (fold_rev (fn th0 => fn (j, accum) => |
502 |
let val th = transfer th0 in |
|
503 |
(j - 1, |
|
504 |
if not (member Thm.eq_thm_prop add_ths th) andalso |
|
505 |
(is_likely_tautology_too_meta_or_too_technical th orelse |
|
506 |
is_too_complex (Thm.prop_of th)) then |
|
507 |
accum |
|
508 |
else |
|
509 |
let |
|
510 |
fun get_name () = |
|
511 |
if name0 = "" orelse name0 = local_thisN then |
|
512 |
backquote_thm ctxt th |
|
513 |
else |
|
514 |
let val short_name = Facts.extern ctxt facts name0 in |
|
515 |
if check_thms short_name then |
|
516 |
short_name |
|
517 |
else |
|
518 |
let val long_name = Name_Space.extern ctxt full_space name0 in |
|
519 |
if check_thms long_name then |
|
520 |
long_name |
|
521 |
else |
|
522 |
name0 |
|
523 |
end |
|
524 |
end |
|
525 |
|> make_name keywords multi j |
|
526 |
val stature = stature_of_thm global assms chained css name0 th |
|
527 |
val new = ((get_name, stature), th) |
|
528 |
in |
|
529 |
(if multi then apsnd else apfst) (cons new) accum |
|
530 |
end) |
|
531 |
end) ths (n, accum)) |
|
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
532 |
end) |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
533 |
in |
| 57729 | 534 |
(* The single-theorem names go before the multiple-theorem ones (e.g., "xxx" vs. "xxx(3)"), so |
535 |
that single names are preferred when both are available. *) |
|
536 |
`I [] |
|
537 |
|> add_facts false fold local_facts (unnamed_locals @ named_locals) |
|
538 |
|> add_facts true Facts.fold_static global_facts global_facts |
|
539 |
|> op @ |
|
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
540 |
end |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
541 |
|
| 58919 | 542 |
fun nearly_all_facts ctxt ho_atp {add, del, only} keywords css chained 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
|
543 |
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
|
544 |
[] |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
545 |
else |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
546 |
let |
| 57729 | 547 |
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
|
548 |
in |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
549 |
(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
|
550 |
maps (map (fn ((name, stature), th) => ((K name, stature), th)) |
| 58919 | 551 |
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
|
552 |
else |
|
54080
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
553 |
let |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58928
diff
changeset
|
554 |
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
|
555 |
val facts = |
| 58919 | 556 |
all_facts ctxt false ho_atp keywords add chained css |
|
54143
18def1c73c79
make sure add: doesn't add duplicates, and works for [no_atp] facts
blanchet
parents:
54142
diff
changeset
|
557 |
|> filter_out ((member Thm.eq_thm_prop del orf |
| 57963 | 558 |
(Named_Theorems.member ctxt @{named_theorems no_atp} andf
|
559 |
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
|
560 |
in |
|
540835cf11ed
more gracefully handle huge theories in relevance filters
blanchet
parents:
54078
diff
changeset
|
561 |
facts |
| 48292 | 562 |
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
|
563 |
|> maybe_instantiate_inducts ctxt hyp_ts concl_t |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
564 |
end |
|
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
565 |
|
|
54142
5f3c1b445253
no fact subsumption -- this only confuses later code, e.g. 'add:'
blanchet
parents:
54112
diff
changeset
|
566 |
fun drop_duplicate_facts facts = |
|
54112
9c0f464d1854
drop only real duplicates, not subsumed facts -- this confuses MaSh
blanchet
parents:
54092
diff
changeset
|
567 |
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
|
568 |
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
|
569 |
end |
|
9c0f464d1854
drop only real duplicates, not subsumed facts -- this confuses MaSh
blanchet
parents:
54092
diff
changeset
|
570 |
|
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff
changeset
|
571 |
end; |