src/HOL/Tools/Sledgehammer/sledgehammer_filter_iter.ML
author blanchet
Wed, 18 Jul 2012 08:44:03 +0200
changeset 49303 255c6e1fd505
parent 49265 1065c307fafe
child 49304 6b65f1ad0e4b
permissions -rw-r--r--
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet@49263
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter_iter.ML
blanchet@38261
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
blanchet@36393
     3
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@40139
     4
blanchet@49263
     5
Sledgehammer's iterative relevance filter.
wenzelm@33317
     6
*)
paulson@15452
     7
blanchet@49263
     8
signature SLEDGEHAMMER_FILTER_ITER =
wenzelm@16802
     9
sig
blanchet@47168
    10
  type stature = ATP_Problem_Generate.stature
blanchet@49265
    11
  type relevance_override = Sledgehammer_Fact.relevance_override
blanchet@49303
    12
  type params = Sledgehammer_Provers.params
blanchet@49303
    13
  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
blanchet@40251
    14
blanchet@43517
    15
  val trace : bool Config.T
blanchet@49234
    16
  val pseudo_abs_name : string
blanchet@49234
    17
  val pseudo_skolem_prefix : string
blanchet@44217
    18
  val const_names_in_fact :
blanchet@44217
    19
    theory -> (string * typ -> term list -> bool * term list) -> term
blanchet@44217
    20
    -> string list
blanchet@49263
    21
  val iterative_relevant_facts :
blanchet@49303
    22
    Proof.context -> params -> string -> int -> relevance_fudge option
blanchet@41314
    23
    -> relevance_override -> thm list -> term list -> term
blanchet@47168
    24
    -> (((unit -> string) * stature) * thm) list
blanchet@47168
    25
    -> ((string * stature) * thm) list
paulson@15347
    26
end;
paulson@15347
    27
blanchet@49263
    28
structure Sledgehammer_Filter_Iter : SLEDGEHAMMER_FILTER_ITER =
paulson@15347
    29
struct
paulson@15347
    30
blanchet@47148
    31
open ATP_Problem_Generate
blanchet@49265
    32
open Sledgehammer_Fact
blanchet@49303
    33
open Sledgehammer_Provers
blanchet@38890
    34
blanchet@43517
    35
val trace =
blanchet@43517
    36
  Attrib.setup_config_bool @{binding sledgehammer_filter_trace} (K false)
blanchet@43517
    37
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@35826
    38
blanchet@37616
    39
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
blanchet@49234
    40
val pseudo_abs_name = sledgehammer_prefix ^ "abs"
blanchet@49234
    41
val pseudo_skolem_prefix = sledgehammer_prefix ^ "sko"
blanchet@39236
    42
val theory_const_suffix = Long_Name.separator ^ " 1"
blanchet@37616
    43
blanchet@46944
    44
fun order_of_type (Type (@{type_name fun}, [T1, T2])) =
blanchet@39165
    45
    Int.max (order_of_type T1 + 1, order_of_type T2)
blanchet@39165
    46
  | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
blanchet@39165
    47
  | order_of_type _ = 0
blanchet@39165
    48
blanchet@39058
    49
(* An abstraction of Isabelle types and first-order terms *)
blanchet@39058
    50
datatype pattern = PVar | PApp of string * pattern list
blanchet@39165
    51
datatype ptype = PType of int * pattern list
blanchet@38983
    52
blanchet@39058
    53
fun string_for_pattern PVar = "_"
blanchet@39058
    54
  | string_for_pattern (PApp (s, ps)) =
blanchet@39058
    55
    if null ps then s else s ^ string_for_patterns ps
blanchet@39058
    56
and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
blanchet@39165
    57
fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
paulson@24287
    58
paulson@24287
    59
(*Is the second type an instance of the first one?*)
blanchet@39059
    60
fun match_pattern (PVar, _) = true
blanchet@39059
    61
  | match_pattern (PApp _, PVar) = false
blanchet@39059
    62
  | match_pattern (PApp (s, ps), PApp (t, qs)) =
blanchet@39059
    63
    s = t andalso match_patterns (ps, qs)
blanchet@39059
    64
and match_patterns (_, []) = true
blanchet@39059
    65
  | match_patterns ([], _) = false
blanchet@39059
    66
  | match_patterns (p :: ps, q :: qs) =
blanchet@39059
    67
    match_pattern (p, q) andalso match_patterns (ps, qs)
blanchet@39165
    68
fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
paulson@24287
    69
blanchet@39058
    70
(* Is there a unifiable constant? *)
blanchet@39062
    71
fun pconst_mem f consts (s, ps) =
blanchet@39165
    72
  exists (curry (match_ptype o f) ps)
blanchet@39062
    73
         (map snd (filter (curry (op =) s o fst) consts))
blanchet@39062
    74
fun pconst_hyper_mem f const_tab (s, ps) =
blanchet@39165
    75
  exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
blanchet@37505
    76
blanchet@39165
    77
fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
blanchet@39165
    78
  | pattern_for_type (TFree (s, _)) = PApp (s, [])
blanchet@39165
    79
  | pattern_for_type (TVar _) = PVar
blanchet@39062
    80
blanchet@38983
    81
(* Pairs a constant with the list of its type instantiations. *)
blanchet@41452
    82
fun ptype thy const x =
blanchet@39165
    83
  (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
blanchet@41452
    84
   else [])
blanchet@41452
    85
fun rich_ptype thy const (s, T) =
blanchet@41452
    86
  PType (order_of_type T, ptype thy const (s, T))
blanchet@41452
    87
fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T))
paulson@24287
    88
blanchet@39165
    89
fun string_for_hyper_pconst (s, ps) =
blanchet@39165
    90
  s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
paulson@24287
    91
blanchet@39058
    92
(* Add a pconstant to the table, but a [] entry means a standard
blanchet@39054
    93
   connective, which we ignore.*)
blanchet@41314
    94
fun add_pconst_to_table also_skolem (s, p) =
blanchet@49234
    95
  if (not also_skolem andalso String.isPrefix pseudo_skolem_prefix s) then I
blanchet@41314
    96
  else Symtab.map_default (s, [p]) (insert (op =) p)
blanchet@39054
    97
blanchet@48948
    98
(* Set constants tend to pull in too many irrelevant facts. We limit the damage
blanchet@48948
    99
   by treating them more or less as if they were built-in but add their
blanchet@48948
   100
   axiomatization at the end. *)
blanchet@48948
   101
val set_consts = [@{const_name Collect}, @{const_name Set.member}]
blanchet@48954
   102
val set_thms = @{thms Collect_mem_eq mem_Collect_eq Collect_cong}
blanchet@48948
   103
blanchet@43597
   104
fun add_pconsts_in_term thy is_built_in_const also_skolems pos =
blanchet@37505
   105
  let
blanchet@39054
   106
    val flip = Option.map not
blanchet@38810
   107
    (* We include free variables, as well as constants, to handle locales. For
blanchet@41453
   108
       each quantifiers that must necessarily be skolemized by the automatic
blanchet@41453
   109
       prover, we introduce a fresh constant to simulate the effect of
blanchet@41453
   110
       Skolemization. *)
blanchet@48948
   111
    fun do_const const ext_arg (x as (s, _)) ts =
blanchet@41584
   112
      let val (built_in, ts) = is_built_in_const x ts in
blanchet@48948
   113
        if member (op =) set_consts s then
blanchet@48948
   114
          fold (do_term ext_arg) ts
blanchet@48948
   115
        else
blanchet@48948
   116
          (not built_in
blanchet@48948
   117
           ? add_pconst_to_table also_skolems (rich_pconst thy const x))
blanchet@48948
   118
          #> fold (do_term false) ts
blanchet@41584
   119
      end
blanchet@43606
   120
    and do_term ext_arg t =
blanchet@39062
   121
      case strip_comb t of
blanchet@48948
   122
        (Const x, ts) => do_const true ext_arg x ts
blanchet@48948
   123
      | (Free x, ts) => do_const false ext_arg x ts
blanchet@39165
   124
      | (Abs (_, T, t'), ts) =>
blanchet@43606
   125
        ((null ts andalso not ext_arg)
blanchet@43606
   126
         (* Since lambdas on the right-hand side of equalities are usually
blanchet@48968
   127
            extensionalized later by "abs_extensionalize_term", we don't
blanchet@48968
   128
            penalize them here. *)
blanchet@49234
   129
         ? add_pconst_to_table true (pseudo_abs_name,
blanchet@49234
   130
                                     PType (order_of_type T + 1, [])))
blanchet@43600
   131
        #> fold (do_term false) (t' :: ts)
blanchet@43600
   132
      | (_, ts) => fold (do_term false) ts
blanchet@39165
   133
    fun do_quantifier will_surely_be_skolemized abs_T body_t =
blanchet@37533
   134
      do_formula pos body_t
blanchet@38986
   135
      #> (if also_skolems andalso will_surely_be_skolemized then
blanchet@49234
   136
            add_pconst_to_table true (pseudo_skolem_prefix ^ serial_string (),
blanchet@49234
   137
                                      PType (order_of_type abs_T, []))
blanchet@38810
   138
          else
blanchet@38810
   139
            I)
blanchet@43606
   140
    and do_term_or_formula ext_arg T =
blanchet@43606
   141
      if T = HOLogic.boolT then do_formula NONE else do_term ext_arg
blanchet@37533
   142
    and do_formula pos t =
blanchet@37533
   143
      case t of
blanchet@39165
   144
        Const (@{const_name all}, _) $ Abs (_, T, t') =>
blanchet@39165
   145
        do_quantifier (pos = SOME false) T t'
blanchet@37533
   146
      | @{const "==>"} $ t1 $ t2 =>
blanchet@37533
   147
        do_formula (flip pos) t1 #> do_formula pos t2
blanchet@37533
   148
      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
blanchet@43606
   149
        do_term_or_formula false T t1 #> do_term_or_formula true T t2
blanchet@37533
   150
      | @{const Trueprop} $ t1 => do_formula pos t1
blanchet@41388
   151
      | @{const False} => I
blanchet@41388
   152
      | @{const True} => I
blanchet@37533
   153
      | @{const Not} $ t1 => do_formula (flip pos) t1
blanchet@39165
   154
      | Const (@{const_name All}, _) $ Abs (_, T, t') =>
blanchet@39165
   155
        do_quantifier (pos = SOME false) T t'
blanchet@39165
   156
      | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
blanchet@39165
   157
        do_quantifier (pos = SOME true) T t'
haftmann@39028
   158
      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
haftmann@39028
   159
      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
haftmann@39019
   160
      | @{const HOL.implies} $ t1 $ t2 =>
blanchet@37533
   161
        do_formula (flip pos) t1 #> do_formula pos t2
haftmann@39093
   162
      | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
blanchet@43606
   163
        do_term_or_formula false T t1 #> do_term_or_formula true T t2
blanchet@38810
   164
      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
blanchet@38810
   165
        $ t1 $ t2 $ t3 =>
blanchet@43600
   166
        do_formula NONE t1 #> fold (do_term_or_formula false T) [t2, t3]
blanchet@39165
   167
      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
blanchet@39165
   168
        do_quantifier (is_some pos) T t'
blanchet@39165
   169
      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
blanchet@39165
   170
        do_quantifier (pos = SOME false) T
blanchet@39165
   171
                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
blanchet@39165
   172
      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
blanchet@39165
   173
        do_quantifier (pos = SOME true) T
blanchet@39165
   174
                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
blanchet@37533
   175
      | (t0 as Const (_, @{typ bool})) $ t1 =>
blanchet@43600
   176
        do_term false t0 #> do_formula pos t1  (* theory constant *)
blanchet@43600
   177
      | _ => do_term false t
blanchet@43597
   178
  in do_formula pos end
paulson@24287
   179
blanchet@49242
   180
fun pconsts_in_fact thy is_built_in_const t =
blanchet@49242
   181
  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
blanchet@49242
   182
              (Symtab.empty |> add_pconsts_in_term thy is_built_in_const true
blanchet@49242
   183
                                                   (SOME true) t) []
blanchet@49242
   184
blanchet@49242
   185
val const_names_in_fact = map fst ooo pconsts_in_fact
blanchet@49242
   186
blanchet@49242
   187
(* Inserts a dummy "constant" referring to the theory name, so that relevance
blanchet@49242
   188
   takes the given theory into account. *)
blanchet@41448
   189
fun theory_constify ({theory_const_rel_weight, theory_const_irrel_weight, ...}
blanchet@41448
   190
                     : relevance_fudge) thy_name t =
blanchet@40251
   191
  if exists (curry (op <) 0.0) [theory_const_rel_weight,
blanchet@40251
   192
                                theory_const_irrel_weight] then
blanchet@41448
   193
    Const (thy_name ^ theory_const_suffix, @{typ bool}) $ t
blanchet@37505
   194
  else
blanchet@41448
   195
    t
blanchet@41448
   196
blanchet@41448
   197
fun theory_const_prop_of fudge th =
blanchet@41448
   198
  theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
blanchet@37505
   199
blanchet@49242
   200
fun pair_consts_fact thy is_built_in_const fudge fact =
blanchet@49242
   201
  case fact |> snd |> theory_const_prop_of fudge
blanchet@49242
   202
            |> pconsts_in_fact thy is_built_in_const of
blanchet@49242
   203
    [] => NONE
blanchet@49242
   204
  | consts => SOME ((fact, consts), NONE)
blanchet@49242
   205
blanchet@38982
   206
(* A two-dimensional symbol table counts frequencies of constants. It's keyed
blanchet@38982
   207
   first by constant name and second by its list of type instantiations. For the
blanchet@39058
   208
   latter, we need a linear ordering on "pattern list". *)
blanchet@37505
   209
blanchet@39058
   210
fun pattern_ord p =
blanchet@38982
   211
  case p of
blanchet@38983
   212
    (PVar, PVar) => EQUAL
blanchet@39058
   213
  | (PVar, PApp _) => LESS
blanchet@39058
   214
  | (PApp _, PVar) => GREATER
blanchet@39058
   215
  | (PApp q1, PApp q2) =>
blanchet@39058
   216
    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
blanchet@39165
   217
fun ptype_ord (PType p, PType q) =
blanchet@39165
   218
  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
paulson@24287
   219
blanchet@39165
   220
structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
paulson@24287
   221
blanchet@40445
   222
fun count_fact_consts thy fudge =
blanchet@37503
   223
  let
blanchet@39062
   224
    fun do_const const (s, T) ts =
blanchet@39062
   225
      (* Two-dimensional table update. Constant maps to types maps to count. *)
blanchet@41452
   226
      PType_Tab.map_default (rich_ptype thy const (s, T), 0) (Integer.add 1)
blanchet@39165
   227
      |> Symtab.map_default (s, PType_Tab.empty)
blanchet@39062
   228
      #> fold do_term ts
blanchet@39062
   229
    and do_term t =
blanchet@39062
   230
      case strip_comb t of
blanchet@39062
   231
        (Const x, ts) => do_const true x ts
blanchet@39062
   232
      | (Free x, ts) => do_const false x ts
blanchet@39062
   233
      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
blanchet@39062
   234
      | (_, ts) => fold do_term ts
blanchet@40251
   235
  in do_term o theory_const_prop_of fudge o snd end
paulson@24287
   236
blanchet@39613
   237
fun pow_int _ 0 = 1.0
blanchet@39165
   238
  | pow_int x 1 = x
blanchet@39165
   239
  | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
blanchet@39165
   240
paulson@24287
   241
(*The frequency of a constant is the sum of those of all instances of its type.*)
blanchet@39059
   242
fun pconst_freq match const_tab (c, ps) =
blanchet@39165
   243
  PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
blanchet@39165
   244
                 (the (Symtab.lookup const_tab c)) 0
blanchet@38925
   245
paulson@24287
   246
blanchet@38331
   247
(* A surprising number of theorems contain only a few significant constants.
blanchet@38331
   248
   These include all induction rules, and other general theorems. *)
blanchet@37503
   249
blanchet@37503
   250
(* "log" seems best in practice. A constant function of one ignores the constant
blanchet@39164
   251
   frequencies. Rare constants give more points if they are relevant than less
blanchet@39164
   252
   rare ones. *)
blanchet@39613
   253
fun rel_weight_for _ freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
blanchet@39164
   254
blanchet@39164
   255
(* Irrelevant constants are treated differently. We associate lower penalties to
blanchet@39164
   256
   very rare constants and very common ones -- the former because they can't
blanchet@39164
   257
   lead to the inclusion of too many new facts, and the latter because they are
blanchet@39164
   258
   so common as to be of little interest. *)
blanchet@40251
   259
fun irrel_weight_for ({worse_irrel_freq, higher_order_irrel_weight, ...}
blanchet@40251
   260
                      : relevance_fudge) order freq =
blanchet@40251
   261
  let val (k, x) = worse_irrel_freq |> `Real.ceil in
blanchet@39165
   262
    (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
blanchet@39165
   263
     else rel_weight_for order freq / rel_weight_for order k)
blanchet@40251
   264
    * pow_int higher_order_irrel_weight (order - 1)
blanchet@39164
   265
  end
blanchet@37503
   266
blanchet@42661
   267
fun multiplier_for_const_name local_const_multiplier s =
blanchet@42661
   268
  if String.isSubstring "." s then 1.0 else local_const_multiplier
blanchet@42661
   269
blanchet@37503
   270
(* Computes a constant's weight, as determined by its frequency. *)
blanchet@42661
   271
fun generic_pconst_weight local_const_multiplier abs_weight skolem_weight
blanchet@43600
   272
                          theory_const_weight chained_const_weight weight_for f
blanchet@43600
   273
                          const_tab chained_const_tab (c as (s, PType (m, _))) =
blanchet@49234
   274
  if s = pseudo_abs_name then
blanchet@42661
   275
    abs_weight
blanchet@49234
   276
  else if String.isPrefix pseudo_skolem_prefix s then
blanchet@42661
   277
    skolem_weight
blanchet@42661
   278
  else if String.isSuffix theory_const_suffix s then
blanchet@42661
   279
    theory_const_weight
blanchet@42661
   280
  else
blanchet@42661
   281
    multiplier_for_const_name local_const_multiplier s
blanchet@42661
   282
    * weight_for m (pconst_freq (match_ptype o f) const_tab c)
blanchet@43600
   283
    |> (if chained_const_weight < 1.0 andalso
blanchet@43600
   284
           pconst_hyper_mem I chained_const_tab c then
blanchet@43600
   285
          curry (op *) chained_const_weight
blanchet@43600
   286
        else
blanchet@43600
   287
          I)
blanchet@39056
   288
blanchet@42661
   289
fun rel_pconst_weight ({local_const_multiplier, abs_rel_weight,
blanchet@42661
   290
                        theory_const_rel_weight, ...} : relevance_fudge)
blanchet@42661
   291
                      const_tab =
blanchet@42661
   292
  generic_pconst_weight local_const_multiplier abs_rel_weight 0.0
blanchet@43600
   293
                        theory_const_rel_weight 0.0 rel_weight_for I const_tab
blanchet@43600
   294
                        Symtab.empty
blanchet@43600
   295
blanchet@42661
   296
fun irrel_pconst_weight (fudge as {local_const_multiplier, abs_irrel_weight,
blanchet@42661
   297
                                   skolem_irrel_weight,
blanchet@43600
   298
                                   theory_const_irrel_weight,
blanchet@43600
   299
                                   chained_const_irrel_weight, ...})
blanchet@43600
   300
                        const_tab chained_const_tab =
blanchet@42661
   301
  generic_pconst_weight local_const_multiplier abs_irrel_weight
blanchet@42661
   302
                        skolem_irrel_weight theory_const_irrel_weight
blanchet@43600
   303
                        chained_const_irrel_weight (irrel_weight_for fudge) swap
blanchet@43600
   304
                        const_tab chained_const_tab
paulson@24287
   305
blanchet@47168
   306
fun stature_bonus ({intro_bonus, ...} : relevance_fudge) (_, Intro) =
blanchet@47168
   307
    intro_bonus
blanchet@47168
   308
  | stature_bonus {elim_bonus, ...} (_, Elim) = elim_bonus
blanchet@47168
   309
  | stature_bonus {simp_bonus, ...} (_, Simp) = simp_bonus
blanchet@47168
   310
  | stature_bonus {local_bonus, ...} (Local, _) = local_bonus
blanchet@47168
   311
  | stature_bonus {assum_bonus, ...} (Assum, _) = assum_bonus
blanchet@47168
   312
  | stature_bonus {chained_bonus, ...} (Chained, _) = chained_bonus
blanchet@47168
   313
  | stature_bonus _ _ = 0.0
blanchet@38990
   314
blanchet@40645
   315
fun is_odd_const_name s =
blanchet@49234
   316
  s = pseudo_abs_name orelse String.isPrefix pseudo_skolem_prefix s orelse
blanchet@40645
   317
  String.isSuffix theory_const_suffix s
blanchet@40645
   318
blanchet@47168
   319
fun fact_weight fudge stature const_tab relevant_consts chained_consts
blanchet@47168
   320
                fact_consts =
blanchet@40445
   321
  case fact_consts |> List.partition (pconst_hyper_mem I relevant_consts)
blanchet@40445
   322
                   ||> filter_out (pconst_hyper_mem swap relevant_consts) of
blanchet@39062
   323
    ([], _) => 0.0
blanchet@38983
   324
  | (rel, irrel) =>
blanchet@40645
   325
    if forall (forall (is_odd_const_name o fst)) [rel, irrel] then
blanchet@40617
   326
      0.0
blanchet@40617
   327
    else
blanchet@40617
   328
      let
blanchet@43600
   329
        val irrel = irrel |> filter_out (pconst_mem swap rel)
blanchet@40617
   330
        val rel_weight =
blanchet@40617
   331
          0.0 |> fold (curry (op +) o rel_pconst_weight fudge const_tab) rel
blanchet@40617
   332
        val irrel_weight =
blanchet@47168
   333
          ~ (stature_bonus fudge stature)
blanchet@43600
   334
          |> fold (curry (op +)
blanchet@43600
   335
                   o irrel_pconst_weight fudge const_tab chained_consts) irrel
blanchet@40617
   336
        val res = rel_weight / (rel_weight + irrel_weight)
blanchet@40617
   337
      in if Real.isFinite res then res else 0.0 end
blanchet@38986
   338
blanchet@38938
   339
type annotated_thm =
blanchet@47168
   340
  (((unit -> string) * stature) * thm) * (string * ptype) list
blanchet@37505
   341
blanchet@43517
   342
fun take_most_relevant ctxt max_relevant remaining_max
blanchet@43593
   343
        ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
blanchet@40251
   344
        (candidates : (annotated_thm * real) list) =
blanchet@38983
   345
  let
blanchet@38986
   346
    val max_imperfect =
blanchet@40251
   347
      Real.ceil (Math.pow (max_imperfect,
blanchet@39130
   348
                    Math.pow (Real.fromInt remaining_max
blanchet@40251
   349
                              / Real.fromInt max_relevant, max_imperfect_exp)))
blanchet@38986
   350
    val (perfect, imperfect) =
blanchet@39115
   351
      candidates |> sort (Real.compare o swap o pairself snd)
blanchet@39115
   352
                 |> take_prefix (fn (_, w) => w > 0.99999)
blanchet@38986
   353
    val ((accepts, more_rejects), rejects) =
blanchet@38986
   354
      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
blanchet@38983
   355
  in
blanchet@43517
   356
    trace_msg ctxt (fn () =>
wenzelm@41739
   357
        "Actually passed (" ^ string_of_int (length accepts) ^ " of " ^
wenzelm@41739
   358
        string_of_int (length candidates) ^ "): " ^
blanchet@39115
   359
        (accepts |> map (fn ((((name, _), _), _), weight) =>
blanchet@38991
   360
                            name () ^ " [" ^ Real.toString weight ^ "]")
blanchet@38984
   361
                 |> commas));
blanchet@38986
   362
    (accepts, more_rejects @ rejects)
blanchet@38983
   363
  end
paulson@24287
   364
blanchet@47168
   365
fun if_empty_replace_with_scope thy is_built_in_const facts sc tab =
blanchet@39054
   366
  if Symtab.is_empty tab then
blanchet@43597
   367
    Symtab.empty
blanchet@43597
   368
    |> fold (add_pconsts_in_term thy is_built_in_const false (SOME false))
blanchet@47168
   369
            (map_filter (fn ((_, (sc', _)), th) =>
blanchet@47168
   370
                            if sc' = sc then SOME (prop_of th) else NONE) facts)
blanchet@39054
   371
  else
blanchet@39054
   372
    tab
blanchet@39054
   373
blanchet@43572
   374
fun consider_arities is_built_in_const th =
blanchet@41406
   375
  let
blanchet@41406
   376
    fun aux _ _ NONE = NONE
blanchet@41406
   377
      | aux t args (SOME tab) =
blanchet@41406
   378
        case t of
blanchet@41406
   379
          t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 []
blanchet@41406
   380
        | Const (x as (s, _)) =>
blanchet@41584
   381
          (if is_built_in_const x args |> fst then
blanchet@41406
   382
             SOME tab
blanchet@41406
   383
           else case Symtab.lookup tab s of
blanchet@41406
   384
             NONE => SOME (Symtab.update (s, length args) tab)
blanchet@41406
   385
           | SOME n => if n = length args then SOME tab else NONE)
blanchet@41406
   386
        | _ => SOME tab
blanchet@41406
   387
  in aux (prop_of th) [] end
blanchet@41406
   388
blanchet@45653
   389
(* FIXME: This is currently only useful for polymorphic type encodings. *)
blanchet@43572
   390
fun could_benefit_from_ext is_built_in_const facts =
blanchet@43572
   391
  fold (consider_arities is_built_in_const o snd) facts (SOME Symtab.empty)
blanchet@41406
   392
  |> is_none
blanchet@41406
   393
blanchet@44363
   394
(* High enough so that it isn't wrongly considered as very relevant (e.g., for E
blanchet@44363
   395
   weights), but low enough so that it is unlikely to be truncated away if few
blanchet@44363
   396
   facts are included. *)
blanchet@44363
   397
val special_fact_index = 75
blanchet@44363
   398
blanchet@49303
   399
fun relevance_filter ctxt thres0 decay max_relevant is_built_in_const
blanchet@40251
   400
        (fudge as {threshold_divisor, ridiculous_threshold, ...})
blanchet@49303
   401
        ({del, ...} : relevance_override) facts chained_ts hyp_ts concl_t =
blanchet@38978
   402
  let
wenzelm@43232
   403
    val thy = Proof_Context.theory_of ctxt
blanchet@40445
   404
    val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
blanchet@43597
   405
    val add_pconsts = add_pconsts_in_term thy is_built_in_const false o SOME
blanchet@43597
   406
    val chained_const_tab = Symtab.empty |> fold (add_pconsts true) chained_ts
blanchet@39054
   407
    val goal_const_tab =
blanchet@43597
   408
      Symtab.empty |> fold (add_pconsts true) hyp_ts
blanchet@43597
   409
                   |> add_pconsts false concl_t
blanchet@43597
   410
      |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab)
blanchet@47168
   411
      |> fold (if_empty_replace_with_scope thy is_built_in_const facts)
blanchet@39237
   412
              [Chained, Assum, Local]
blanchet@39256
   413
    val del_ths = Attrib.eval_thms ctxt del
blanchet@43798
   414
    val facts = facts |> filter_out (member Thm.eq_thm_prop del_ths o snd)
blanchet@49303
   415
    fun iter j remaining_max thres rel_const_tab hopeless hopeful =
blanchet@38978
   416
      let
blanchet@40432
   417
        fun relevant [] _ [] =
blanchet@38986
   418
            (* Nothing has been added this iteration. *)
blanchet@49303
   419
            if j = 0 andalso thres >= ridiculous_threshold then
blanchet@38986
   420
              (* First iteration? Try again. *)
blanchet@49303
   421
              iter 0 max_relevant (thres / threshold_divisor) rel_const_tab
blanchet@38986
   422
                   hopeless hopeful
blanchet@38983
   423
            else
blanchet@40432
   424
              []
blanchet@39115
   425
          | relevant candidates rejects [] =
blanchet@38978
   426
            let
blanchet@38986
   427
              val (accepts, more_rejects) =
blanchet@43517
   428
                take_most_relevant ctxt max_relevant remaining_max fudge
blanchet@43517
   429
                                   candidates
blanchet@38978
   430
              val rel_const_tab' =
blanchet@38984
   431
                rel_const_tab
blanchet@41314
   432
                |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
blanchet@38983
   433
              fun is_dirty (c, _) =
blanchet@38983
   434
                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
blanchet@38984
   435
              val (hopeful_rejects, hopeless_rejects) =
blanchet@38984
   436
                 (rejects @ hopeless, ([], []))
blanchet@38984
   437
                 |-> fold (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   438
                              if exists is_dirty consts then
blanchet@38984
   439
                                apfst (cons (ax, NONE))
blanchet@38984
   440
                              else
blanchet@38984
   441
                                apsnd (cons (ax, old_weight)))
blanchet@38984
   442
                 |>> append (more_rejects
blanchet@38984
   443
                             |> map (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   444
                                        (ax, if exists is_dirty consts then NONE
blanchet@38984
   445
                                             else SOME old_weight)))
blanchet@49303
   446
              val thres =
blanchet@49303
   447
                1.0 - (1.0 - thres)
blanchet@39057
   448
                      * Math.pow (decay, Real.fromInt (length accepts))
blanchet@38986
   449
              val remaining_max = remaining_max - length accepts
blanchet@38978
   450
            in
blanchet@43517
   451
              trace_msg ctxt (fn () => "New or updated constants: " ^
blanchet@38983
   452
                  commas (rel_const_tab' |> Symtab.dest
blanchet@39057
   453
                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
blanchet@39062
   454
                          |> map string_for_hyper_pconst));
blanchet@38984
   455
              map (fst o fst) accepts @
blanchet@38986
   456
              (if remaining_max = 0 then
blanchet@40432
   457
                 []
blanchet@38984
   458
               else
blanchet@49303
   459
                 iter (j + 1) remaining_max thres rel_const_tab'
blanchet@38986
   460
                      hopeless_rejects hopeful_rejects)
blanchet@38978
   461
            end
blanchet@39115
   462
          | relevant candidates rejects
blanchet@47168
   463
                     (((ax as (((_, stature), _), fact_consts)), cached_weight)
blanchet@38986
   464
                      :: hopeful) =
blanchet@38978
   465
            let
blanchet@38978
   466
              val weight =
blanchet@38978
   467
                case cached_weight of
blanchet@38978
   468
                  SOME w => w
blanchet@47168
   469
                | NONE => fact_weight fudge stature const_tab rel_const_tab
blanchet@43597
   470
                                      chained_const_tab fact_consts
blanchet@38978
   471
            in
blanchet@49303
   472
              if weight >= thres then
blanchet@39115
   473
                relevant ((ax, weight) :: candidates) rejects hopeful
blanchet@38978
   474
              else
blanchet@39115
   475
                relevant candidates ((ax, weight) :: rejects) hopeful
blanchet@38978
   476
            end
blanchet@38978
   477
        in
blanchet@43517
   478
          trace_msg ctxt (fn () =>
blanchet@38983
   479
              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
blanchet@49303
   480
              Real.toString thres ^ ", constants: " ^
blanchet@38983
   481
              commas (rel_const_tab |> Symtab.dest
blanchet@38983
   482
                      |> filter (curry (op <>) [] o snd)
blanchet@39062
   483
                      |> map string_for_hyper_pconst));
blanchet@39115
   484
          relevant [] [] hopeful
blanchet@38978
   485
        end
blanchet@48948
   486
    fun uses_const s t =
blanchet@48948
   487
      fold_aterms (curry (fn (Const (s', _), false) => s' = s | (_, b) => b)) t
blanchet@48948
   488
                  false
blanchet@48948
   489
    fun uses_const_anywhere accepts s =
blanchet@48948
   490
      exists (uses_const s o prop_of o snd) accepts orelse
blanchet@48948
   491
      exists (uses_const s) (concl_t :: hyp_ts)
blanchet@48948
   492
    fun add_set_const_thms accepts =
blanchet@48948
   493
      exists (uses_const_anywhere accepts) set_consts ? append set_thms
blanchet@44363
   494
    fun insert_into_facts accepts [] = accepts
blanchet@44363
   495
      | insert_into_facts accepts ths =
blanchet@44363
   496
        let
blanchet@44363
   497
          val add = facts |> filter (member Thm.eq_thm_prop ths o snd)
blanchet@44363
   498
          val (bef, after) =
blanchet@44363
   499
            accepts |> filter_out (member Thm.eq_thm_prop ths o snd)
blanchet@44363
   500
                    |> take (max_relevant - length add)
blanchet@44363
   501
                    |> chop special_fact_index
blanchet@44363
   502
        in bef @ add @ after end
blanchet@44363
   503
    fun insert_special_facts accepts =
blanchet@48954
   504
       (* FIXME: get rid of "ext" here once it is treated as a helper *)
blanchet@43907
   505
       [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
blanchet@48948
   506
          |> add_set_const_thms accepts
blanchet@44363
   507
          |> insert_into_facts accepts
blanchet@38978
   508
  in
blanchet@40615
   509
    facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
blanchet@49303
   510
          |> iter 0 max_relevant thres0 goal_const_tab []
blanchet@44363
   511
          |> insert_special_facts
blanchet@43517
   512
          |> tap (fn accepts => trace_msg ctxt (fn () =>
wenzelm@41739
   513
                      "Total relevant: " ^ string_of_int (length accepts)))
blanchet@38978
   514
  end
paulson@24287
   515
blanchet@49303
   516
fun iterative_relevant_facts ctxt
blanchet@49303
   517
        ({relevance_thresholds = (thres0, thres1), ...} : params) prover
blanchet@49303
   518
        max_relevant fudge override chained_ths hyp_ts concl_t facts =
blanchet@37534
   519
  let
wenzelm@43232
   520
    val thy = Proof_Context.theory_of ctxt
blanchet@49303
   521
    val is_built_in_const =
blanchet@49303
   522
      Sledgehammer_Provers.is_built_in_const_for_prover ctxt prover
blanchet@49303
   523
    val fudge =
blanchet@49303
   524
      case fudge of
blanchet@49303
   525
        SOME fudge => fudge
blanchet@49303
   526
      | NONE => Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover
blanchet@49303
   527
    val decay = Math.pow ((1.0 - thres1) / (1.0 - thres0),
blanchet@39057
   528
                          1.0 / Real.fromInt (max_relevant + 1))
blanchet@37534
   529
  in
blanchet@43517
   530
    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^
blanchet@43517
   531
                             " facts");
blanchet@49303
   532
    (if thres1 < 0.0 then
blanchet@40445
   533
       facts
blanchet@49303
   534
     else if thres0 > 1.0 orelse thres0 > thres1 then
blanchet@38978
   535
       []
blanchet@38978
   536
     else
blanchet@49303
   537
       relevance_filter ctxt thres0 decay max_relevant is_built_in_const
blanchet@45483
   538
           fudge override facts (chained_ths |> map prop_of) hyp_ts
blanchet@45483
   539
           (concl_t |> theory_constify fudge (Context.theory_name thy)))
blanchet@39057
   540
    |> map (apfst (apfst (fn f => f ())))
blanchet@37534
   541
  end
immler@30536
   542
paulson@15347
   543
end;