| author | wenzelm | 
| Fri, 20 Mar 2015 22:18:40 +0100 | |
| changeset 59771 | c6e60787ffe2 | 
| parent 59582 | 0fbed69ff081 | 
| child 60638 | 16d80e5ef2dc | 
| permissions | -rw-r--r-- | 
| 48380 | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_mepo.ML | 
| 38027 | 2 | Author: Jia Meng, Cambridge University Computer Laboratory and NICTA | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36227diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 39958 | 4 | |
| 48380 | 5 | Sledgehammer's iterative relevance filter (MePo = Meng-Paulson). | 
| 33309 | 6 | *) | 
| 15452 | 7 | |
| 48381 | 8 | signature SLEDGEHAMMER_MEPO = | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 9 | sig | 
| 46340 | 10 | type stature = ATP_Problem_Generate.stature | 
| 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: 
50985diff
changeset | 11 | type raw_fact = Sledgehammer_Fact.raw_fact | 
| 48296 
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
 blanchet parents: 
48293diff
changeset | 12 | type fact = Sledgehammer_Fact.fact | 
| 55201 | 13 | type params = Sledgehammer_Prover.params | 
| 54095 | 14 | |
| 15 | type relevance_fudge = | |
| 16 |     {local_const_multiplier : real,
 | |
| 17 | worse_irrel_freq : real, | |
| 18 | higher_order_irrel_weight : real, | |
| 19 | abs_rel_weight : real, | |
| 20 | abs_irrel_weight : real, | |
| 21 | theory_const_rel_weight : real, | |
| 22 | theory_const_irrel_weight : real, | |
| 23 | chained_const_irrel_weight : real, | |
| 24 | intro_bonus : real, | |
| 25 | elim_bonus : real, | |
| 26 | simp_bonus : real, | |
| 27 | local_bonus : real, | |
| 28 | assum_bonus : real, | |
| 29 | chained_bonus : real, | |
| 30 | max_imperfect : real, | |
| 31 | max_imperfect_exp : real, | |
| 32 | threshold_divisor : real, | |
| 33 | ridiculous_threshold : real} | |
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
39958diff
changeset | 34 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42641diff
changeset | 35 | val trace : bool Config.T | 
| 48219 | 36 | val pseudo_abs_name : string | 
| 54095 | 37 | val default_relevance_fudge : relevance_fudge | 
| 57149 | 38 | val mepo_suggested_facts : Proof.context -> params -> int -> relevance_fudge option -> | 
| 39 | term list -> term -> raw_fact list -> fact list | |
| 15347 | 40 | end; | 
| 41 | ||
| 48381 | 42 | structure Sledgehammer_MePo : SLEDGEHAMMER_MEPO = | 
| 15347 | 43 | struct | 
| 44 | ||
| 46320 | 45 | open ATP_Problem_Generate | 
| 50608 | 46 | open Sledgehammer_Util | 
| 48296 
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
 blanchet parents: 
48293diff
changeset | 47 | open Sledgehammer_Fact | 
| 55201 | 48 | open Sledgehammer_Prover | 
| 38652 
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
 blanchet parents: 
38644diff
changeset | 49 | |
| 57149 | 50 | val trace = Attrib.setup_config_bool @{binding sledgehammer_mepo_trace} (K false)
 | 
| 51 | ||
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42641diff
changeset | 52 | fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () | 
| 35826 | 53 | |
| 37616 
c8d2d84d6011
always perform relevance filtering on original formulas
 blanchet parents: 
37580diff
changeset | 54 | val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator | 
| 48219 | 55 | val pseudo_abs_name = sledgehammer_prefix ^ "abs" | 
| 38992 
542474156c66
introduce fudge factors to deal with "theory const"
 blanchet parents: 
38988diff
changeset | 56 | val theory_const_suffix = Long_Name.separator ^ " 1" | 
| 37616 
c8d2d84d6011
always perform relevance filtering on original formulas
 blanchet parents: 
37580diff
changeset | 57 | |
| 54095 | 58 | type relevance_fudge = | 
| 59 |   {local_const_multiplier : real,
 | |
| 60 | worse_irrel_freq : real, | |
| 61 | higher_order_irrel_weight : real, | |
| 62 | abs_rel_weight : real, | |
| 63 | abs_irrel_weight : real, | |
| 64 | theory_const_rel_weight : real, | |
| 65 | theory_const_irrel_weight : real, | |
| 66 | chained_const_irrel_weight : real, | |
| 67 | intro_bonus : real, | |
| 68 | elim_bonus : real, | |
| 69 | simp_bonus : real, | |
| 70 | local_bonus : real, | |
| 71 | assum_bonus : real, | |
| 72 | chained_bonus : real, | |
| 73 | max_imperfect : real, | |
| 74 | max_imperfect_exp : real, | |
| 75 | threshold_divisor : real, | |
| 76 | ridiculous_threshold : real} | |
| 77 | ||
| 78 | (* FUDGE *) | |
| 79 | val default_relevance_fudge = | |
| 80 |   {local_const_multiplier = 1.5,
 | |
| 81 | worse_irrel_freq = 100.0, | |
| 82 | higher_order_irrel_weight = 1.05, | |
| 83 | abs_rel_weight = 0.5, | |
| 84 | abs_irrel_weight = 2.0, | |
| 85 | theory_const_rel_weight = 0.5, | |
| 86 | theory_const_irrel_weight = 0.25, | |
| 87 | chained_const_irrel_weight = 0.25, | |
| 88 | intro_bonus = 0.15, | |
| 89 | elim_bonus = 0.15, | |
| 90 | simp_bonus = 0.15, | |
| 91 | local_bonus = 0.55, | |
| 92 | assum_bonus = 1.05, | |
| 93 | chained_bonus = 1.5, | |
| 94 | max_imperfect = 11.5, | |
| 95 | max_imperfect_exp = 1.0, | |
| 96 | threshold_divisor = 2.0, | |
| 97 | ridiculous_threshold = 0.1} | |
| 98 | ||
| 46073 
b2594cc862d7
removed special handling for set constants in relevance filter
 blanchet parents: 
45982diff
changeset | 99 | fun order_of_type (Type (@{type_name fun}, [T1, T2])) =
 | 
| 38939 | 100 | Int.max (order_of_type T1 + 1, order_of_type T2) | 
| 101 | | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0 | |
| 102 | | order_of_type _ = 0 | |
| 103 | ||
| 38823 
828e68441a2f
renaming + treat "TFree" better in "pattern_for_type"
 blanchet parents: 
38822diff
changeset | 104 | (* An abstraction of Isabelle types and first-order terms *) | 
| 
828e68441a2f
renaming + treat "TFree" better in "pattern_for_type"
 blanchet parents: 
38822diff
changeset | 105 | datatype pattern = PVar | PApp of string * pattern list | 
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 106 | datatype ptype = PType of int * typ list | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 107 | |
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 108 | fun string_of_patternT (TVar _) = "_" | 
| 57149 | 109 | | string_of_patternT (Type (s, ps)) = if null ps then s else s ^ string_of_patternsT ps | 
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 110 | | string_of_patternT (TFree (s, _)) = s | 
| 54086 | 111 | and string_of_patternsT ps = "(" ^ commas (map string_of_patternT ps) ^ ")"
 | 
| 112 | fun string_of_ptype (PType (_, ps)) = string_of_patternsT ps | |
| 24287 | 113 | |
| 114 | (*Is the second type an instance of the first one?*) | |
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 115 | fun match_patternT (TVar _, _) = true | 
| 57149 | 116 | | match_patternT (Type (s, ps), Type (t, qs)) = s = t andalso match_patternsT (ps, qs) | 
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 117 | | match_patternT (TFree (s, _), TFree (t, _)) = s = t | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 118 | | match_patternT (_, _) = false | 
| 54086 | 119 | and match_patternsT (_, []) = true | 
| 120 | | match_patternsT ([], _) = false | |
| 57149 | 121 | | match_patternsT (p :: ps, q :: qs) = match_patternT (p, q) andalso match_patternsT (ps, qs) | 
| 54086 | 122 | fun match_ptype (PType (_, ps), PType (_, qs)) = match_patternsT (ps, qs) | 
| 24287 | 123 | |
| 38823 
828e68441a2f
renaming + treat "TFree" better in "pattern_for_type"
 blanchet parents: 
38822diff
changeset | 124 | (* Is there a unifiable constant? *) | 
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 125 | fun pconst_mem f consts (s, ps) = | 
| 57149 | 126 | exists (curry (match_ptype o f) ps) (map snd (filter (curry (op =) s o fst) consts)) | 
| 127 | ||
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 128 | fun pconst_hyper_mem f const_tab (s, ps) = | 
| 38939 | 129 | exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s)) | 
| 37505 
d9af5c01dc4a
added code to optionally perform fact filtering on the original (non-CNF) formulas
 blanchet parents: 
37504diff
changeset | 130 | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 131 | (* Pairs a constant with the list of its type instantiations. *) | 
| 57149 | 132 | fun ptype thy const x = (if const then these (try (Sign.const_typargs thy) x) else []) | 
| 133 | fun rich_ptype thy const (s, T) = PType (order_of_type T, ptype thy const (s, T)) | |
| 41204 
bd57cf5944cb
get rid of experimental feature of term patterns in relevance filter -- doesn't work well unless we take into consideration the equality theory entailed by the relevant facts
 blanchet parents: 
41202diff
changeset | 134 | fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T)) | 
| 24287 | 135 | |
| 57149 | 136 | fun string_of_hyper_pconst (s, ps) = s ^ "{" ^ commas (map string_of_ptype ps) ^ "}"
 | 
| 24287 | 137 | |
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 138 | fun patternT_eq (TVar _, TVar _) = true | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 139 | | patternT_eq (Type (s, Ts), Type (t, Us)) = s = t andalso patternsT_eq (Ts, Us) | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 140 | | patternT_eq (TFree (s, _), TFree (t, _)) = (s = t) | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 141 | | patternT_eq _ = false | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 142 | and patternsT_eq ([], []) = true | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 143 | | patternsT_eq ([], _) = false | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 144 | | patternsT_eq (_, []) = false | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 145 | | patternsT_eq (T :: Ts, U :: Us) = patternT_eq (T, U) andalso patternsT_eq (Ts, Us) | 
| 57149 | 146 | |
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 147 | fun ptype_eq (PType (m, Ts), PType (n, Us)) = m = n andalso patternsT_eq (Ts, Us) | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 148 | |
| 57149 | 149 | (* Add a pconstant to the table, but a [] entry means a standard connective, which we ignore. *) | 
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 150 | fun add_pconst_to_table (s, p) = Symtab.map_default (s, [p]) (insert ptype_eq p) | 
| 38819 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 151 | |
| 57149 | 152 | (* Set constants tend to pull in too many irrelevant facts. We limit the damage by treating them | 
| 153 | more or less as if they were built-in but add their axiomatization at the end. *) | |
| 47933 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 154 | val set_consts = [@{const_name Collect}, @{const_name Set.member}]
 | 
| 47939 
9ff976a6c2cb
added "Collect_cong" to cover extensionality of "Collect" (special cases of "ext" pass through the relevant filter)
 blanchet parents: 
47933diff
changeset | 155 | val set_thms = @{thms Collect_mem_eq mem_Collect_eq Collect_cong}
 | 
| 47933 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 156 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 157 | fun add_pconsts_in_term thy = | 
| 37505 
d9af5c01dc4a
added code to optionally perform fact filtering on the original (non-CNF) formulas
 blanchet parents: 
37504diff
changeset | 158 | let | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 159 | fun do_const const (x as (s, _)) ts = | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 160 | if member (op =) set_consts s then | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 161 | fold (do_term false) ts | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 162 | else | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 163 | (not (is_irrelevant_const s) ? add_pconst_to_table (rich_pconst thy const x)) | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 164 | #> fold (do_term false) ts | 
| 42741 
546b0bda3cb8
do not pollute relevance filter facts with too many facts about the boring set constants Collect and mem_def, which we might anyway unfold depending on Meson's settings
 blanchet parents: 
42738diff
changeset | 165 | and do_term ext_arg t = | 
| 55286 | 166 | (case strip_comb t of | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 167 | (Const x, ts) => do_const true x ts | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 168 | | (Free x, ts) => do_const false x ts | 
| 38939 | 169 | | (Abs (_, T, t'), ts) => | 
| 42741 
546b0bda3cb8
do not pollute relevance filter facts with too many facts about the boring set constants Collect and mem_def, which we might anyway unfold depending on Meson's settings
 blanchet parents: 
42738diff
changeset | 170 | ((null ts andalso not ext_arg) | 
| 57149 | 171 | (* Since lambdas on the right-hand side of equalities are usually extensionalized later by | 
| 172 | "abs_extensionalize_term", we don't penalize them here. *) | |
| 173 | ? add_pconst_to_table (pseudo_abs_name, PType (order_of_type T + 1, []))) | |
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 174 | #> fold (do_term false) (t' :: ts) | 
| 55286 | 175 | | (_, ts) => fold (do_term false) ts) | 
| 42741 
546b0bda3cb8
do not pollute relevance filter facts with too many facts about the boring set constants Collect and mem_def, which we might anyway unfold depending on Meson's settings
 blanchet parents: 
42738diff
changeset | 176 | and do_term_or_formula ext_arg T = | 
| 53551 | 177 | if T = HOLogic.boolT then do_formula else do_term ext_arg | 
| 178 | and do_formula t = | |
| 55286 | 179 | (case t of | 
| 56245 | 180 |         Const (@{const_name Pure.all}, _) $ Abs (_, _, t') => do_formula t'
 | 
| 181 |       | @{const Pure.imp} $ t1 $ t2 => do_formula t1 #> do_formula t2
 | |
| 182 |       | Const (@{const_name Pure.eq}, Type (_, [T, _])) $ t1 $ t2 =>
 | |
| 42741 
546b0bda3cb8
do not pollute relevance filter facts with too many facts about the boring set constants Collect and mem_def, which we might anyway unfold depending on Meson's settings
 blanchet parents: 
42738diff
changeset | 183 | do_term_or_formula false T t1 #> do_term_or_formula true T t2 | 
| 53551 | 184 |       | @{const Trueprop} $ t1 => do_formula t1
 | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41138diff
changeset | 185 |       | @{const False} => I
 | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41138diff
changeset | 186 |       | @{const True} => I
 | 
| 53551 | 187 |       | @{const Not} $ t1 => do_formula t1
 | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 188 |       | Const (@{const_name All}, _) $ Abs (_, _, t') => do_formula t'
 | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 189 |       | Const (@{const_name Ex}, _) $ Abs (_, _, t') => do_formula t'
 | 
| 53551 | 190 |       | @{const HOL.conj} $ t1 $ t2 => do_formula t1 #> do_formula t2
 | 
| 191 |       | @{const HOL.disj} $ t1 $ t2 => do_formula t1 #> do_formula t2
 | |
| 192 |       | @{const HOL.implies} $ t1 $ t2 => do_formula t1 #> do_formula t2
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38829diff
changeset | 193 |       | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
 | 
| 42741 
546b0bda3cb8
do not pollute relevance filter facts with too many facts about the boring set constants Collect and mem_def, which we might anyway unfold depending on Meson's settings
 blanchet parents: 
42738diff
changeset | 194 | do_term_or_formula false T t1 #> do_term_or_formula true T t2 | 
| 57149 | 195 |       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])])) $ t1 $ t2 $ t3 =>
 | 
| 53551 | 196 | do_formula t1 #> fold (do_term_or_formula false T) [t2, t3] | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 197 |       | Const (@{const_name Ex1}, _) $ Abs (_, _, t') => do_formula t'
 | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 198 |       | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, t') =>
 | 
| 53551 | 199 | do_formula (t1 $ Bound ~1) #> do_formula t' | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 200 |       | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, t') =>
 | 
| 53551 | 201 | do_formula (t1 $ Bound ~1) #> do_formula t' | 
| 37537 | 202 |       | (t0 as Const (_, @{typ bool})) $ t1 =>
 | 
| 53551 | 203 | do_term false t0 #> do_formula t1 (* theory constant *) | 
| 55286 | 204 | | _ => do_term false t) | 
| 57149 | 205 | in | 
| 206 | do_formula | |
| 207 | end | |
| 24287 | 208 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 209 | fun pconsts_in_fact thy t = | 
| 57149 | 210 | Symtab.fold (fn (s, pss) => fold (cons o pair s) pss) (Symtab.empty |> add_pconsts_in_term thy t) | 
| 211 | [] | |
| 48227 | 212 | |
| 213 | (* Inserts a dummy "constant" referring to the theory name, so that relevance | |
| 214 | takes the given theory into account. *) | |
| 57149 | 215 | fun theory_constify ({theory_const_rel_weight, theory_const_irrel_weight, ...} : relevance_fudge)
 | 
| 216 | thy_name t = | |
| 217 | if exists (curry (op <) 0.0) [theory_const_rel_weight, theory_const_irrel_weight] then | |
| 41200 
6cc9b6fd7f6f
add the current theory's constant to the goal to make theorems from the current theory more relevant on the first iteration already
 blanchet parents: 
41199diff
changeset | 218 |     Const (thy_name ^ theory_const_suffix, @{typ bool}) $ t
 | 
| 37505 
d9af5c01dc4a
added code to optionally perform fact filtering on the original (non-CNF) formulas
 blanchet parents: 
37504diff
changeset | 219 | else | 
| 41200 
6cc9b6fd7f6f
add the current theory's constant to the goal to make theorems from the current theory more relevant on the first iteration already
 blanchet parents: 
41199diff
changeset | 220 | t | 
| 
6cc9b6fd7f6f
add the current theory's constant to the goal to make theorems from the current theory more relevant on the first iteration already
 blanchet parents: 
41199diff
changeset | 221 | |
| 
6cc9b6fd7f6f
add the current theory's constant to the goal to make theorems from the current theory more relevant on the first iteration already
 blanchet parents: 
41199diff
changeset | 222 | fun theory_const_prop_of fudge th = | 
| 59582 | 223 | theory_constify fudge (Context.theory_name (Thm.theory_of_thm th)) (Thm.prop_of th) | 
| 37505 
d9af5c01dc4a
added code to optionally perform fact filtering on the original (non-CNF) formulas
 blanchet parents: 
37504diff
changeset | 224 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 225 | fun pair_consts_fact thy fudge fact = | 
| 55286 | 226 | (case fact |> snd |> theory_const_prop_of fudge |> pconsts_in_fact thy of | 
| 48227 | 227 | [] => NONE | 
| 55286 | 228 | | consts => SOME ((fact, consts), NONE)) | 
| 48227 | 229 | |
| 38743 | 230 | (* A two-dimensional symbol table counts frequencies of constants. It's keyed | 
| 231 | first by constant name and second by its list of type instantiations. For the | |
| 38823 
828e68441a2f
renaming + treat "TFree" better in "pattern_for_type"
 blanchet parents: 
38822diff
changeset | 232 | latter, we need a linear ordering on "pattern list". *) | 
| 24287 | 233 | |
| 54086 | 234 | fun patternT_ord p = | 
| 55286 | 235 | (case p of | 
| 54088 | 236 | (Type (s, ps), Type (t, qs)) => | 
| 237 | (case fast_string_ord (s, t) of | |
| 238 | EQUAL => dict_ord patternT_ord (ps, qs) | |
| 239 | | ord => ord) | |
| 240 | | (TVar _, TVar _) => EQUAL | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 241 | | (TVar _, _) => LESS | 
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 242 | | (Type _, TVar _) => GREATER | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 243 | | (Type _, TFree _) => LESS | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 244 | | (TFree (s, _), TFree (t, _)) => fast_string_ord (s, t) | 
| 55286 | 245 | | (TFree _, _) => GREATER) | 
| 57149 | 246 | |
| 54088 | 247 | fun ptype_ord (PType (m, ps), PType (n, qs)) = | 
| 248 | (case dict_ord patternT_ord (ps, qs) of | |
| 249 | EQUAL => int_ord (m, n) | |
| 250 | | ord => ord) | |
| 24287 | 251 | |
| 38939 | 252 | structure PType_Tab = Table(type key = ptype val ord = ptype_ord) | 
| 24287 | 253 | |
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40191diff
changeset | 254 | fun count_fact_consts thy fudge = | 
| 37503 
c2dfa26b9da6
cosmetics + prevent consideration of inlined Skolem terms in relevance filter
 blanchet parents: 
37502diff
changeset | 255 | let | 
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 256 | fun do_const const (s, T) ts = | 
| 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 257 | (* Two-dimensional table update. Constant maps to types maps to count. *) | 
| 41204 
bd57cf5944cb
get rid of experimental feature of term patterns in relevance filter -- doesn't work well unless we take into consideration the equality theory entailed by the relevant facts
 blanchet parents: 
41202diff
changeset | 258 | PType_Tab.map_default (rich_ptype thy const (s, T), 0) (Integer.add 1) | 
| 38939 | 259 | |> Symtab.map_default (s, PType_Tab.empty) | 
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 260 | #> fold do_term ts | 
| 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 261 | and do_term t = | 
| 55286 | 262 | (case strip_comb t of | 
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 263 | (Const x, ts) => do_const true x ts | 
| 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 264 | | (Free x, ts) => do_const false x ts | 
| 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 265 | | (Abs (_, _, t'), ts) => fold do_term (t' :: ts) | 
| 55286 | 266 | | (_, ts) => fold do_term ts) | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
39958diff
changeset | 267 | in do_term o theory_const_prop_of fudge o snd end | 
| 24287 | 268 | |
| 39367 | 269 | fun pow_int _ 0 = 1.0 | 
| 38939 | 270 | | pow_int x 1 = x | 
| 271 | | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x | |
| 272 | ||
| 24287 | 273 | (*The frequency of a constant is the sum of those of all instances of its type.*) | 
| 38824 | 274 | fun pconst_freq match const_tab (c, ps) = | 
| 57149 | 275 | PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m) (the (Symtab.lookup const_tab c)) 0 | 
| 38686 | 276 | |
| 57149 | 277 | (* A surprising number of theorems contain only a few significant constants. These include all | 
| 278 | induction rules and other general theorems. *) | |
| 37503 
c2dfa26b9da6
cosmetics + prevent consideration of inlined Skolem terms in relevance filter
 blanchet parents: 
37502diff
changeset | 279 | |
| 
c2dfa26b9da6
cosmetics + prevent consideration of inlined Skolem terms in relevance filter
 blanchet parents: 
37502diff
changeset | 280 | (* "log" seems best in practice. A constant function of one ignores the constant | 
| 38938 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 281 | frequencies. Rare constants give more points if they are relevant than less | 
| 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 282 | rare ones. *) | 
| 39367 | 283 | fun rel_weight_for _ freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0) | 
| 38938 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 284 | |
| 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 285 | (* Irrelevant constants are treated differently. We associate lower penalties to | 
| 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 286 | very rare constants and very common ones -- the former because they can't | 
| 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 287 | lead to the inclusion of too many new facts, and the latter because they are | 
| 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 288 | so common as to be of little interest. *) | 
| 57149 | 289 | fun irrel_weight_for ({worse_irrel_freq, higher_order_irrel_weight, ...} : relevance_fudge) order
 | 
| 290 | freq = | |
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
39958diff
changeset | 291 | let val (k, x) = worse_irrel_freq |> `Real.ceil in | 
| 38939 | 292 | (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x | 
| 293 | else rel_weight_for order freq / rel_weight_for order k) | |
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
39958diff
changeset | 294 | * pow_int higher_order_irrel_weight (order - 1) | 
| 38938 
2b93dbc07778
improve weighting of irrelevant constants, based on Mirabelle experiments
 blanchet parents: 
38937diff
changeset | 295 | end | 
| 37503 
c2dfa26b9da6
cosmetics + prevent consideration of inlined Skolem terms in relevance filter
 blanchet parents: 
37502diff
changeset | 296 | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51026diff
changeset | 297 | fun multiplier_of_const_name local_const_multiplier s = | 
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 298 | if String.isSubstring "." s then 1.0 else local_const_multiplier | 
| 38821 | 299 | |
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 300 | (* Computes a constant's weight, as determined by its frequency. *) | 
| 57149 | 301 | fun generic_pconst_weight local_const_multiplier abs_weight theory_const_weight chained_const_weight | 
| 302 | weight_for f const_tab chained_const_tab (c as (s, PType (m, _))) = | |
| 48219 | 303 | if s = pseudo_abs_name then | 
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 304 | abs_weight | 
| 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 305 | else if String.isSuffix theory_const_suffix s then | 
| 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 306 | theory_const_weight | 
| 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 307 | else | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51026diff
changeset | 308 | multiplier_of_const_name local_const_multiplier s | 
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 309 | * weight_for m (pconst_freq (match_ptype o f) const_tab c) | 
| 57149 | 310 | |> (if chained_const_weight < 1.0 andalso pconst_hyper_mem I chained_const_tab c then | 
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 311 | curry (op *) chained_const_weight | 
| 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 312 | else | 
| 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 313 | I) | 
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 314 | |
| 57149 | 315 | fun rel_pconst_weight ({local_const_multiplier, abs_rel_weight, theory_const_rel_weight,
 | 
| 316 | ...} : relevance_fudge) const_tab = | |
| 317 | generic_pconst_weight local_const_multiplier abs_rel_weight theory_const_rel_weight 0.0 | |
| 318 | rel_weight_for I const_tab Symtab.empty | |
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 319 | |
| 41790 
56dcd46ddf7a
give more weight to Frees than to Consts in relevance filter
 blanchet parents: 
41768diff
changeset | 320 | fun irrel_pconst_weight (fudge as {local_const_multiplier, abs_irrel_weight,
 | 
| 57149 | 321 | theory_const_irrel_weight, chained_const_irrel_weight, ...}) const_tab chained_const_tab = | 
| 322 | generic_pconst_weight local_const_multiplier abs_irrel_weight theory_const_irrel_weight | |
| 323 | chained_const_irrel_weight (irrel_weight_for fudge) swap const_tab chained_const_tab | |
| 24287 | 324 | |
| 57149 | 325 | fun stature_bonus ({intro_bonus, ...} : relevance_fudge) (_, Intro) = intro_bonus
 | 
| 46340 | 326 |   | stature_bonus {elim_bonus, ...} (_, Elim) = elim_bonus
 | 
| 327 |   | stature_bonus {simp_bonus, ...} (_, Simp) = simp_bonus
 | |
| 328 |   | stature_bonus {local_bonus, ...} (Local, _) = local_bonus
 | |
| 329 |   | stature_bonus {assum_bonus, ...} (Assum, _) = assum_bonus
 | |
| 330 |   | stature_bonus {chained_bonus, ...} (Chained, _) = chained_bonus
 | |
| 331 | | stature_bonus _ _ = 0.0 | |
| 38751 
01c4d14b2a61
add a bonus for chained facts, since they are likely to be relevant;
 blanchet parents: 
38749diff
changeset | 332 | |
| 40418 | 333 | fun is_odd_const_name s = | 
| 53550 
ffe2ce7910c1
get rid of another complication in relevance filter
 blanchet parents: 
53548diff
changeset | 334 | s = pseudo_abs_name orelse String.isSuffix theory_const_suffix s | 
| 40418 | 335 | |
| 53516 
925591242376
got rid of currently unused data structure, to speed up relevance filter
 blanchet parents: 
53493diff
changeset | 336 | fun fact_weight fudge stature const_tab rel_const_tab chained_const_tab | 
| 
925591242376
got rid of currently unused data structure, to speed up relevance filter
 blanchet parents: 
53493diff
changeset | 337 | fact_consts = | 
| 55286 | 338 | (case fact_consts |> List.partition (pconst_hyper_mem I rel_const_tab) | 
| 53158 
4b9df3461eda
thread information about when a constant became relevant through MePo -- the information is not used yet but could help fight starvation in the (ITP) world
 blanchet parents: 
53127diff
changeset | 339 | ||> filter_out (pconst_hyper_mem swap rel_const_tab) of | 
| 38827 
cf01645cbbce
extended relevance filter with first-order term matching
 blanchet parents: 
38825diff
changeset | 340 | ([], _) => 0.0 | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 341 | | (rel, irrel) => | 
| 40418 | 342 | if forall (forall (is_odd_const_name o fst)) [rel, irrel] then | 
| 40371 | 343 | 0.0 | 
| 344 | else | |
| 345 | let | |
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42732diff
changeset | 346 | val irrel = irrel |> filter_out (pconst_mem swap rel) | 
| 57149 | 347 | val rel_weight = 0.0 |> fold (curry (op +) o rel_pconst_weight fudge const_tab) rel | 
| 40371 | 348 | val irrel_weight = | 
| 46340 | 349 | ~ (stature_bonus fudge stature) | 
| 57149 | 350 | |> fold (curry (op +) o irrel_pconst_weight fudge const_tab chained_const_tab) irrel | 
| 40371 | 351 | val res = rel_weight / (rel_weight + irrel_weight) | 
| 57149 | 352 | in | 
| 353 | if Real.isFinite res then res else 0.0 | |
| 354 | end) | |
| 38747 | 355 | |
| 48293 | 356 | fun take_most_relevant ctxt max_facts remaining_max | 
| 57149 | 357 |     ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
 | 
| 358 | (candidates : ((raw_fact * (string * ptype) list) * real) list) = | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 359 | let | 
| 38747 | 360 | val max_imperfect = | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
39958diff
changeset | 361 | Real.ceil (Math.pow (max_imperfect, | 
| 57149 | 362 | Math.pow (Real.fromInt remaining_max / Real.fromInt max_facts, max_imperfect_exp))) | 
| 363 | val (perfect, imperfect) = candidates | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57149diff
changeset | 364 | |> sort (Real.compare o swap o apply2 snd) | 
| 57149 | 365 | |> take_prefix (fn (_, w) => w > 0.99999) | 
| 38747 | 366 | val ((accepts, more_rejects), rejects) = | 
| 367 | chop max_imperfect imperfect |>> append perfect |>> chop remaining_max | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 368 | in | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42641diff
changeset | 369 | trace_msg ctxt (fn () => | 
| 57149 | 370 |       "Actually passed (" ^ string_of_int (length accepts) ^ " of " ^
 | 
| 371 | string_of_int (length candidates) ^ "): " ^ | |
| 372 | (accepts | |
| 373 | |> map (fn ((((name, _), _), _), weight) => name () ^ " [" ^ Real.toString weight ^ "]") | |
| 374 | |> commas)); | |
| 38747 | 375 | (accepts, more_rejects @ rejects) | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 376 | end | 
| 24287 | 377 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 378 | fun if_empty_replace_with_scope thy facts sc tab = | 
| 38819 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 379 | if Symtab.is_empty tab then | 
| 42732 
86683865278d
no penality for constants that appear in chained facts
 blanchet parents: 
42730diff
changeset | 380 | Symtab.empty | 
| 57149 | 381 | |> fold (add_pconsts_in_term thy) (map_filter (fn ((_, (sc', _)), th) => | 
| 59582 | 382 | if sc' = sc then SOME (Thm.prop_of th) else NONE) facts) | 
| 38819 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 383 | else | 
| 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 384 | tab | 
| 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 385 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 386 | fun consider_arities th = | 
| 41158 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 387 | let | 
| 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 388 | fun aux _ _ NONE = NONE | 
| 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 389 | | aux t args (SOME tab) = | 
| 55286 | 390 | (case t of | 
| 41158 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 391 | t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 [] | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 392 | | Const (s, _) => | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 393 | (if is_widely_irrelevant_const s then | 
| 41158 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 394 | SOME tab | 
| 55286 | 395 | else | 
| 396 | (case Symtab.lookup tab s of | |
| 397 | NONE => SOME (Symtab.update (s, length args) tab) | |
| 398 | | SOME n => if n = length args then SOME tab else NONE)) | |
| 399 | | _ => SOME tab) | |
| 57149 | 400 | in | 
| 59582 | 401 | aux (Thm.prop_of th) [] | 
| 57149 | 402 | end | 
| 41158 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 403 | |
| 44785 | 404 | (* FIXME: This is currently only useful for polymorphic type encodings. *) | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 405 | fun could_benefit_from_ext facts = | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 406 | fold (consider_arities o snd) facts (SOME Symtab.empty) |> is_none | 
| 41158 
8c9c31a757f5
make Sledgehammer's relevance filter include the "ext" rule when appropriate
 blanchet parents: 
41140diff
changeset | 407 | |
| 43492 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 408 | (* High enough so that it isn't wrongly considered as very relevant (e.g., for E | 
| 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 409 | weights), but low enough so that it is unlikely to be truncated away if few | 
| 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 410 | facts are included. *) | 
| 53546 
a2d2fa096e31
kick out totally hopeless facts after 5 iterations to speed things up
 blanchet parents: 
53516diff
changeset | 411 | val special_fact_index = 45 (* FUDGE *) | 
| 
a2d2fa096e31
kick out totally hopeless facts after 5 iterations to speed things up
 blanchet parents: 
53516diff
changeset | 412 | |
| 54087 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 413 | fun eq_prod eqx eqy ((x1, y1), (x2, y2)) = eqx (x1, x2) andalso eqy (y1, y2) | 
| 
957115f3dae4
use plain types instead of dedicated type pattern
 blanchet parents: 
54086diff
changeset | 414 | |
| 53546 
a2d2fa096e31
kick out totally hopeless facts after 5 iterations to speed things up
 blanchet parents: 
53516diff
changeset | 415 | val really_hopeless_get_kicked_out_iter = 5 (* FUDGE *) | 
| 43492 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 416 | |
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 417 | fun relevance_filter ctxt thres0 decay max_facts | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 418 |         (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts concl_t =
 | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 419 | let | 
| 42361 | 420 | val thy = Proof_Context.theory_of ctxt | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40191diff
changeset | 421 | val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 422 | val add_pconsts = add_pconsts_in_term thy | 
| 48292 | 423 | val chained_ts = | 
| 59582 | 424 | facts |> map_filter (try (fn ((_, (Chained, _)), th) => Thm.prop_of th)) | 
| 53551 | 425 | val chained_const_tab = Symtab.empty |> fold add_pconsts chained_ts | 
| 38819 
71c9f61516cd
if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
 blanchet parents: 
38818diff
changeset | 426 | val goal_const_tab = | 
| 53158 
4b9df3461eda
thread information about when a constant became relevant through MePo -- the information is not used yet but could help fight starvation in the (ITP) world
 blanchet parents: 
53127diff
changeset | 427 | Symtab.empty | 
| 53551 | 428 | |> fold add_pconsts hyp_ts | 
| 429 | |> add_pconsts concl_t | |
| 42732 
86683865278d
no penality for constants that appear in chained facts
 blanchet parents: 
42730diff
changeset | 430 | |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab) | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 431 | |> fold (if_empty_replace_with_scope thy facts) [Chained, Assum, Local] | 
| 57149 | 432 | |
| 53516 
925591242376
got rid of currently unused data structure, to speed up relevance filter
 blanchet parents: 
53493diff
changeset | 433 | fun iter j remaining_max thres rel_const_tab hopeless hopeful = | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 434 | let | 
| 53546 
a2d2fa096e31
kick out totally hopeless facts after 5 iterations to speed things up
 blanchet parents: 
53516diff
changeset | 435 | val hopeless = | 
| 57149 | 436 | hopeless |> j = really_hopeless_get_kicked_out_iter ? filter_out (fn (_, w) => w < 0.001) | 
| 40191 
257d2e06bfb8
put theorems added using "add:" at the beginning of the list returned by the relevance filter, so that they don't get truncated away
 blanchet parents: 
40071diff
changeset | 437 | fun relevant [] _ [] = | 
| 38747 | 438 | (* Nothing has been added this iteration. *) | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 439 | if j = 0 andalso thres >= ridiculous_threshold then | 
| 38747 | 440 | (* First iteration? Try again. *) | 
| 57149 | 441 | iter 0 max_facts (thres / threshold_divisor) rel_const_tab hopeless hopeful | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 442 | else | 
| 40191 
257d2e06bfb8
put theorems added using "add:" at the beginning of the list returned by the relevance filter, so that they don't get truncated away
 blanchet parents: 
40071diff
changeset | 443 | [] | 
| 38889 | 444 | | relevant candidates rejects [] = | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 445 | let | 
| 38747 | 446 | val (accepts, more_rejects) = | 
| 48293 | 447 | take_most_relevant ctxt max_facts remaining_max fudge candidates | 
| 53546 
a2d2fa096e31
kick out totally hopeless facts after 5 iterations to speed things up
 blanchet parents: 
53516diff
changeset | 448 | val sps = maps (snd o fst) accepts | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 449 | val rel_const_tab' = | 
| 53550 
ffe2ce7910c1
get rid of another complication in relevance filter
 blanchet parents: 
53548diff
changeset | 450 | rel_const_tab |> fold add_pconst_to_table sps | 
| 57149 | 451 | |
| 452 | fun is_dirty (s, _) = Symtab.lookup rel_const_tab' s <> Symtab.lookup rel_const_tab s | |
| 453 | ||
| 38745 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 454 | val (hopeful_rejects, hopeless_rejects) = | 
| 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 455 | (rejects @ hopeless, ([], [])) | 
| 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 456 | |-> fold (fn (ax as (_, consts), old_weight) => | 
| 57149 | 457 | if exists is_dirty consts then apfst (cons (ax, NONE)) | 
| 458 | else apsnd (cons (ax, old_weight))) | |
| 38745 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 459 | |>> append (more_rejects | 
| 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 460 | |> map (fn (ax as (_, consts), old_weight) => | 
| 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 461 | (ax, if exists is_dirty consts then NONE | 
| 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 462 | else SOME old_weight))) | 
| 57149 | 463 | val thres = 1.0 - (1.0 - thres) * Math.pow (decay, Real.fromInt (length accepts)) | 
| 38747 | 464 | val remaining_max = remaining_max - length accepts | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 465 | in | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42641diff
changeset | 466 | trace_msg ctxt (fn () => "New or updated constants: " ^ | 
| 57149 | 467 | commas (rel_const_tab' | 
| 468 | |> Symtab.dest | |
| 469 | |> subtract (eq_prod (op =) (eq_list ptype_eq)) (Symtab.dest rel_const_tab) | |
| 470 | |> map string_of_hyper_pconst)); | |
| 38745 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 471 | map (fst o fst) accepts @ | 
| 38747 | 472 | (if remaining_max = 0 then | 
| 40191 
257d2e06bfb8
put theorems added using "add:" at the beginning of the list returned by the relevance filter, so that they don't get truncated away
 blanchet parents: 
40071diff
changeset | 473 | [] | 
| 38745 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 blanchet parents: 
38744diff
changeset | 474 | else | 
| 57149 | 475 | iter (j + 1) remaining_max thres rel_const_tab' hopeless_rejects hopeful_rejects) | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 476 | end | 
| 38889 | 477 | | relevant candidates rejects | 
| 57149 | 478 | (((ax as (((_, stature), _), fact_consts)), cached_weight) :: hopeful) = | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 479 | let | 
| 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 480 | val weight = | 
| 55286 | 481 | (case cached_weight of | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 482 | SOME w => w | 
| 53158 
4b9df3461eda
thread information about when a constant became relevant through MePo -- the information is not used yet but could help fight starvation in the (ITP) world
 blanchet parents: 
53127diff
changeset | 483 | | NONE => | 
| 55286 | 484 | fact_weight fudge stature const_tab rel_const_tab chained_const_tab fact_consts) | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 485 | in | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 486 | if weight >= thres then | 
| 38889 | 487 | relevant ((ax, weight) :: candidates) rejects hopeful | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 488 | else | 
| 38889 | 489 | relevant candidates ((ax, weight) :: rejects) hopeful | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 490 | end | 
| 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 491 | in | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42641diff
changeset | 492 | trace_msg ctxt (fn () => | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 493 | "ITERATION " ^ string_of_int j ^ ": current threshold: " ^ | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 494 | Real.toString thres ^ ", constants: " ^ | 
| 53158 
4b9df3461eda
thread information about when a constant became relevant through MePo -- the information is not used yet but could help fight starvation in the (ITP) world
 blanchet parents: 
53127diff
changeset | 495 | commas (rel_const_tab | 
| 
4b9df3461eda
thread information about when a constant became relevant through MePo -- the information is not used yet but could help fight starvation in the (ITP) world
 blanchet parents: 
53127diff
changeset | 496 | |> Symtab.dest | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38743diff
changeset | 497 | |> filter (curry (op <>) [] o snd) | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51026diff
changeset | 498 | |> map string_of_hyper_pconst)); | 
| 38889 | 499 | relevant [] [] hopeful | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 500 | end | 
| 47933 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 501 | fun uses_const s t = | 
| 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 502 | fold_aterms (curry (fn (Const (s', _), false) => s' = s | (_, b) => b)) t | 
| 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 503 | false | 
| 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 504 | fun uses_const_anywhere accepts s = | 
| 59582 | 505 | exists (uses_const s o Thm.prop_of o snd) accepts orelse | 
| 47933 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 506 | exists (uses_const s) (concl_t :: hyp_ts) | 
| 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 507 | fun add_set_const_thms accepts = | 
| 
4e8e0245e8be
treat sets specially in relevance filter, as they used to, to avoid cluttering the problem with facts about Set.member and Collect
 blanchet parents: 
47904diff
changeset | 508 | exists (uses_const_anywhere accepts) set_consts ? append set_thms | 
| 43492 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 509 | fun insert_into_facts accepts [] = accepts | 
| 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 510 | | insert_into_facts accepts ths = | 
| 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 511 | let | 
| 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 512 | val add = facts |> filter (member Thm.eq_thm_prop ths o snd) | 
| 57149 | 513 | val (bef, after) = accepts | 
| 514 | |> filter_out (member Thm.eq_thm_prop ths o snd) | |
| 515 | |> take (max_facts - length add) | |
| 516 | |> chop special_fact_index | |
| 517 | in | |
| 518 | bef @ add @ after | |
| 519 | end | |
| 43492 
43326cadc31a
insert rather than append special facts to make it less likely that they're truncated away
 blanchet parents: 
43477diff
changeset | 520 | fun insert_special_facts accepts = | 
| 51026 | 521 | (* FIXME: get rid of "ext" here once it is treated as a helper *) | 
| 57149 | 522 | [] | 
| 523 |       |> could_benefit_from_ext accepts ? cons @{thm ext}
 | |
| 524 | |> add_set_const_thms accepts | |
| 525 | |> insert_into_facts accepts | |
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 526 | in | 
| 57149 | 527 | facts | 
| 528 | |> map_filter (pair_consts_fact thy fudge) | |
| 529 | |> iter 0 max_facts thres0 goal_const_tab [] | |
| 530 | |> insert_special_facts | |
| 531 | |> tap (fn accepts => trace_msg ctxt (fn () => | |
| 532 | "Total relevant: " ^ string_of_int (length accepts))) | |
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 533 | end | 
| 24287 | 534 | |
| 54095 | 535 | fun mepo_suggested_facts ctxt ({fact_thresholds = (thres0, thres1), ...} : params) max_facts fudge
 | 
| 57149 | 536 | hyp_ts concl_t facts = | 
| 37538 
97ab019d5ac8
make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
 blanchet parents: 
37537diff
changeset | 537 | let | 
| 42361 | 538 | val thy = Proof_Context.theory_of ctxt | 
| 54095 | 539 | val fudge = fudge |> the_default default_relevance_fudge | 
| 57149 | 540 | val decay = Math.pow ((1.0 - thres1) / (1.0 - thres0), 1.0 / Real.fromInt (max_facts + 1)) | 
| 37538 
97ab019d5ac8
make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
 blanchet parents: 
37537diff
changeset | 541 | in | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 542 | trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^ " facts"); | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 543 | (if thres1 < 0.0 then | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40191diff
changeset | 544 | facts | 
| 54124 | 545 | else if thres0 > 1.0 orelse thres0 > thres1 orelse max_facts <= 0 then | 
| 38739 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 546 | [] | 
| 
8b8ed80b5699
renamed "relevance_convergence" to "relevance_decay"
 blanchet parents: 
38738diff
changeset | 547 | else | 
| 54089 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 548 | relevance_filter ctxt thres0 decay max_facts fudge facts hyp_ts | 
| 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 blanchet parents: 
54088diff
changeset | 549 | (concl_t |> theory_constify fudge (Context.theory_name thy))) | 
| 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: 
50985diff
changeset | 550 | |> map fact_of_raw_fact | 
| 37538 
97ab019d5ac8
make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
 blanchet parents: 
37537diff
changeset | 551 | end | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 552 | |
| 15347 | 553 | end; |