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