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