src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Thu, 26 Aug 2010 00:49:04 +0200
changeset 38986 b264ae66cede
parent 38984 ad577fd62ee4
child 38988 0d2f7f0614d1
permissions -rw-r--r--
fiddle with relevance filter
blanchet@35826
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
blanchet@38261
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
blanchet@36393
     3
    Author:     Jasmin Blanchette, TU Muenchen
wenzelm@33317
     4
*)
paulson@15452
     5
blanchet@35826
     6
signature SLEDGEHAMMER_FACT_FILTER =
wenzelm@16802
     7
sig
blanchet@35964
     8
  type relevance_override =
blanchet@35964
     9
    {add: Facts.ref list,
blanchet@35964
    10
     del: Facts.ref list,
blanchet@35964
    11
     only: bool}
blanchet@35964
    12
blanchet@37616
    13
  val trace : bool Unsynchronized.ref
blanchet@38983
    14
  val name_thm_pairs_from_ref :
blanchet@38935
    15
    Proof.context -> unit Symtab.table -> thm list -> Facts.ref
blanchet@38983
    16
    -> ((unit -> string * bool) * (bool * thm)) list
blanchet@37347
    17
  val relevant_facts :
blanchet@38984
    18
    bool -> real * real -> int -> bool -> relevance_override
blanchet@38229
    19
    -> Proof.context * (thm list * 'a) -> term list -> term
blanchet@38937
    20
    -> ((string * bool) * thm) list
paulson@15347
    21
end;
paulson@15347
    22
blanchet@35826
    23
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
paulson@15347
    24
struct
paulson@15347
    25
blanchet@38890
    26
open Sledgehammer_Util
blanchet@38890
    27
blanchet@37616
    28
val trace = Unsynchronized.ref false
blanchet@37616
    29
fun trace_msg msg = if !trace then tracing (msg ()) else ()
blanchet@35826
    30
blanchet@37580
    31
val respect_no_atp = true
blanchet@37505
    32
blanchet@35964
    33
type relevance_override =
blanchet@35964
    34
  {add: Facts.ref list,
blanchet@35964
    35
   del: Facts.ref list,
blanchet@35964
    36
   only: bool}
paulson@21070
    37
blanchet@37616
    38
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
blanchet@37616
    39
blanchet@38983
    40
fun repair_name reserved multi j name =
blanchet@38983
    41
  (name |> Symtab.defined reserved name ? quote) ^
blanchet@38983
    42
  (if multi then "(" ^ Int.toString j ^ ")" else "")
blanchet@38983
    43
blanchet@38983
    44
fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
blanchet@38983
    45
  let
blanchet@38983
    46
    val ths = ProofContext.get_fact ctxt xref
blanchet@38983
    47
    val name = Facts.string_of_ref xref
blanchet@38983
    48
    val multi = length ths > 1
blanchet@38983
    49
  in
blanchet@38983
    50
    fold (fn th => fn (j, rest) =>
blanchet@38983
    51
             (j + 1, (fn () => (repair_name reserved multi j name,
blanchet@38983
    52
                                member Thm.eq_thm chained_ths th),
blanchet@38983
    53
                      (multi, th)) :: rest))
blanchet@38983
    54
         ths (1, [])
blanchet@38983
    55
    |> snd
blanchet@38938
    56
  end
blanchet@37616
    57
wenzelm@28477
    58
(***************************************************************)
wenzelm@28477
    59
(* Relevance Filtering                                         *)
wenzelm@28477
    60
(***************************************************************)
mengj@19194
    61
paulson@24287
    62
(*** constants with types ***)
paulson@24287
    63
paulson@24287
    64
(*An abstraction of Isabelle types*)
blanchet@38983
    65
datatype pseudotype = PVar | PType of string * pseudotype list
blanchet@38983
    66
blanchet@38983
    67
fun string_for_pseudotype PVar = "?"
blanchet@38983
    68
  | string_for_pseudotype (PType (s, Ts)) =
blanchet@38983
    69
    (case Ts of
blanchet@38983
    70
       [] => ""
blanchet@38983
    71
     | [T] => string_for_pseudotype T
blanchet@38983
    72
     | Ts => string_for_pseudotypes Ts ^ " ") ^ s
blanchet@38983
    73
and string_for_pseudotypes Ts =
blanchet@38983
    74
  "(" ^ commas (map string_for_pseudotype Ts) ^ ")"
paulson@24287
    75
paulson@24287
    76
(*Is the second type an instance of the first one?*)
blanchet@38983
    77
fun match_pseudotype (PType (a, T), PType (b, U)) =
blanchet@38983
    78
    a = b andalso match_pseudotypes (T, U)
blanchet@38983
    79
  | match_pseudotype (PVar, _) = true
blanchet@38983
    80
  | match_pseudotype (_, PVar) = false
blanchet@38983
    81
and match_pseudotypes ([], []) = true
blanchet@38983
    82
  | match_pseudotypes (T :: Ts, U :: Us) =
blanchet@38983
    83
    match_pseudotype (T, U) andalso match_pseudotypes (Ts, Us)
paulson@24287
    84
paulson@24287
    85
(*Is there a unifiable constant?*)
blanchet@38983
    86
fun pseudoconst_mem f const_tab (c, c_typ) =
blanchet@38983
    87
  exists (curry (match_pseudotypes o f) c_typ)
blanchet@38983
    88
         (these (Symtab.lookup const_tab c))
blanchet@37505
    89
blanchet@38983
    90
fun pseudotype_for (Type (c,typs)) = PType (c, map pseudotype_for typs)
blanchet@38983
    91
  | pseudotype_for (TFree _) = PVar
blanchet@38983
    92
  | pseudotype_for (TVar _) = PVar
blanchet@38983
    93
(* Pairs a constant with the list of its type instantiations. *)
blanchet@38983
    94
fun pseudoconst_for thy (c, T) =
blanchet@38983
    95
  (c, map pseudotype_for (Sign.const_typargs thy (c, T)))
blanchet@38983
    96
  handle TYPE _ => (c, [])  (* Variable (locale constant): monomorphic *)
paulson@24287
    97
blanchet@38983
    98
fun string_for_pseudoconst (s, []) = s
blanchet@38983
    99
  | string_for_pseudoconst (s, Ts) = s ^ string_for_pseudotypes Ts
blanchet@38983
   100
fun string_for_super_pseudoconst (s, [[]]) = s
blanchet@38983
   101
  | string_for_super_pseudoconst (s, Tss) =
blanchet@38983
   102
    s ^ "{" ^ commas (map string_for_pseudotypes Tss) ^ "}"
paulson@24287
   103
blanchet@38986
   104
val skolem_prefix = "Sledgehammer."
blanchet@38986
   105
blanchet@38986
   106
(* Add a pseudoconstant to the table, but a [] entry means a standard
blanchet@38986
   107
   connective, which we ignore.*)
blanchet@38986
   108
fun add_pseudoconst_to_table also_skolem (c, ctyps) =
blanchet@38986
   109
  if also_skolem orelse not (String.isPrefix skolem_prefix c) then
blanchet@38986
   110
    Symtab.map_default (c, [ctyps])
blanchet@38986
   111
                       (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
blanchet@38986
   112
  else
blanchet@38986
   113
    I
paulson@24287
   114
blanchet@38931
   115
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
blanchet@38931
   116
blanchet@37533
   117
val flip = Option.map not
blanchet@38337
   118
(* These are typically simplified away by "Meson.presimplify". *)
blanchet@38921
   119
val boring_consts =
blanchet@38921
   120
  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}]
blanchet@37533
   121
blanchet@38986
   122
fun get_pseudoconsts thy also_skolems pos ts =
blanchet@37505
   123
  let
blanchet@38810
   124
    (* We include free variables, as well as constants, to handle locales. For
blanchet@38810
   125
       each quantifiers that must necessarily be skolemized by the ATP, we
blanchet@38810
   126
       introduce a fresh constant to simulate the effect of Skolemization. *)
blanchet@37533
   127
    fun do_term t =
blanchet@37533
   128
      case t of
blanchet@38986
   129
        Const x => add_pseudoconst_to_table also_skolems (pseudoconst_for thy x)
blanchet@38986
   130
      | Free (s, _) => add_pseudoconst_to_table also_skolems (s, [])
blanchet@38927
   131
      | t1 $ t2 => fold do_term [t1, t2]
blanchet@38986
   132
      | Abs (_, _, t') => do_term t'  (* FIXME: add penalty? *)
blanchet@37533
   133
      | _ => I
blanchet@38810
   134
    fun do_quantifier will_surely_be_skolemized body_t =
blanchet@37533
   135
      do_formula pos body_t
blanchet@38986
   136
      #> (if also_skolems andalso will_surely_be_skolemized then
blanchet@38986
   137
            add_pseudoconst_to_table true (gensym skolem_prefix, [])
blanchet@38810
   138
          else
blanchet@38810
   139
            I)
blanchet@38810
   140
    and do_term_or_formula T =
blanchet@38931
   141
      if is_formula_type T then do_formula NONE else do_term
blanchet@37533
   142
    and do_formula pos t =
blanchet@37533
   143
      case t of
blanchet@37533
   144
        Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
blanchet@38810
   145
        do_quantifier (pos = SOME false) body_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@38810
   149
        fold (do_term_or_formula T) [t1, t2]
blanchet@37533
   150
      | @{const Trueprop} $ t1 => do_formula pos t1
blanchet@37533
   151
      | @{const Not} $ t1 => do_formula (flip pos) t1
blanchet@37533
   152
      | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
blanchet@38810
   153
        do_quantifier (pos = SOME false) body_t
blanchet@37533
   154
      | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
blanchet@38810
   155
        do_quantifier (pos = SOME true) body_t
blanchet@37533
   156
      | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
blanchet@37533
   157
      | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
blanchet@37533
   158
      | @{const "op -->"} $ t1 $ t2 =>
blanchet@37533
   159
        do_formula (flip pos) t1 #> do_formula pos t2
blanchet@37533
   160
      | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
blanchet@38810
   161
        fold (do_term_or_formula T) [t1, t2]
blanchet@38810
   162
      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
blanchet@38810
   163
        $ t1 $ t2 $ t3 =>
blanchet@38810
   164
        do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
blanchet@38810
   165
      | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
blanchet@38810
   166
        do_quantifier (is_some pos) body_t
blanchet@38810
   167
      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
blanchet@38810
   168
        do_quantifier (pos = SOME false)
blanchet@38810
   169
                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
blanchet@38810
   170
      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
blanchet@38810
   171
        do_quantifier (pos = SOME true)
blanchet@38810
   172
                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
blanchet@37533
   173
      | (t0 as Const (_, @{typ bool})) $ t1 =>
blanchet@37533
   174
        do_term t0 #> do_formula pos t1  (* theory constant *)
blanchet@37533
   175
      | _ => do_term t
blanchet@37505
   176
  in
blanchet@38337
   177
    Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
blanchet@38229
   178
                 |> fold (do_formula pos) ts
blanchet@37505
   179
  end
paulson@24287
   180
paulson@24287
   181
(*Inserts a dummy "constant" referring to the theory name, so that relevance
paulson@24287
   182
  takes the given theory into account.*)
blanchet@37616
   183
fun theory_const_prop_of theory_relevant th =
blanchet@37505
   184
  if theory_relevant then
blanchet@37505
   185
    let
blanchet@37505
   186
      val name = Context.theory_name (theory_of_thm th)
blanchet@37505
   187
      val t = Const (name ^ ". 1", @{typ bool})
blanchet@37505
   188
    in t $ prop_of th end
blanchet@37505
   189
  else
blanchet@37505
   190
    prop_of th
blanchet@37505
   191
paulson@24287
   192
(**** Constant / Type Frequencies ****)
paulson@24287
   193
blanchet@38982
   194
(* A two-dimensional symbol table counts frequencies of constants. It's keyed
blanchet@38982
   195
   first by constant name and second by its list of type instantiations. For the
blanchet@38983
   196
   latter, we need a linear ordering on "pseudotype list". *)
blanchet@37505
   197
blanchet@38983
   198
fun pseudotype_ord p =
blanchet@38982
   199
  case p of
blanchet@38983
   200
    (PVar, PVar) => EQUAL
blanchet@38983
   201
  | (PVar, PType _) => LESS
blanchet@38983
   202
  | (PType _, PVar) => GREATER
blanchet@38983
   203
  | (PType q1, PType q2) =>
blanchet@38983
   204
    prod_ord fast_string_ord (dict_ord pseudotype_ord) (q1, q2)
paulson@24287
   205
blanchet@38982
   206
structure CTtab =
blanchet@38983
   207
  Table(type key = pseudotype list val ord = dict_ord pseudotype_ord)
paulson@24287
   208
blanchet@37616
   209
fun count_axiom_consts theory_relevant thy (_, th) =
blanchet@37503
   210
  let
blanchet@37503
   211
    fun do_const (a, T) =
blanchet@38983
   212
      let val (c, cts) = pseudoconst_for thy (a, T) in
blanchet@37503
   213
        (* Two-dimensional table update. Constant maps to types maps to
blanchet@37503
   214
           count. *)
blanchet@37503
   215
        CTtab.map_default (cts, 0) (Integer.add 1)
blanchet@37503
   216
        |> Symtab.map_default (c, CTtab.empty)
blanchet@37503
   217
      end
blanchet@37503
   218
    fun do_term (Const x) = do_const x
blanchet@37503
   219
      | do_term (Free x) = do_const x
blanchet@37503
   220
      | do_term (t $ u) = do_term t #> do_term u
blanchet@37503
   221
      | do_term (Abs (_, _, t)) = do_term t
blanchet@37503
   222
      | do_term _ = I
blanchet@37616
   223
  in th |> theory_const_prop_of theory_relevant |> do_term end
paulson@24287
   224
paulson@24287
   225
paulson@24287
   226
(**** Actual Filtering Code ****)
paulson@24287
   227
paulson@24287
   228
(*The frequency of a constant is the sum of those of all instances of its type.*)
blanchet@38983
   229
fun pseudoconst_freq match const_tab (c, cts) =
blanchet@38983
   230
  CTtab.fold (fn (cts', m) => match (cts, cts') ? Integer.add m)
blanchet@38925
   231
             (the (Symtab.lookup const_tab c)) 0
blanchet@38925
   232
  handle Option.Option => 0
blanchet@38925
   233
paulson@24287
   234
blanchet@38331
   235
(* A surprising number of theorems contain only a few significant constants.
blanchet@38331
   236
   These include all induction rules, and other general theorems. *)
blanchet@37503
   237
blanchet@37503
   238
(* "log" seems best in practice. A constant function of one ignores the constant
blanchet@37503
   239
   frequencies. *)
blanchet@38986
   240
fun rel_log n = 1.0 + 2.0 / Math.ln (Real.fromInt n + 1.0)
blanchet@38986
   241
(* TODO: experiment
blanchet@38986
   242
fun irrel_log n = 0.5 + 1.0 / Math.ln (Real.fromInt n + 1.0)
blanchet@38986
   243
*)
blanchet@38986
   244
fun irrel_log n = Math.ln (Real.fromInt n + 19.0) / 6.4
blanchet@37503
   245
blanchet@37503
   246
(* Computes a constant's weight, as determined by its frequency. *)
blanchet@38986
   247
val rel_weight = rel_log oo pseudoconst_freq match_pseudotypes
blanchet@38986
   248
fun irrel_weight const_tab (c as (s, _)) =
blanchet@38986
   249
  if String.isPrefix skolem_prefix s then 1.0
blanchet@38986
   250
  else irrel_log (pseudoconst_freq (match_pseudotypes o swap) const_tab c)
blanchet@38986
   251
(* TODO: experiment
blanchet@38986
   252
fun irrel_weight _ _ = 1.0
blanchet@38986
   253
*)
paulson@24287
   254
blanchet@38926
   255
fun axiom_weight const_tab relevant_consts axiom_consts =
blanchet@38983
   256
  case axiom_consts |> List.partition (pseudoconst_mem I relevant_consts)
blanchet@38983
   257
                    ||> filter_out (pseudoconst_mem swap relevant_consts) of
blanchet@38983
   258
    ([], []) => 0.0
blanchet@38983
   259
  | (_, []) => 1.0
blanchet@38983
   260
  | (rel, irrel) =>
blanchet@38983
   261
    let
blanchet@38983
   262
      val rel_weight = fold (curry Real.+ o rel_weight const_tab) rel 0.0
blanchet@38983
   263
      val irrel_weight = fold (curry Real.+ o irrel_weight const_tab) irrel 0.0
blanchet@38983
   264
      val res = rel_weight / (rel_weight + irrel_weight)
blanchet@38983
   265
    in if Real.isFinite res then res else 0.0 end
blanchet@37505
   266
blanchet@38986
   267
(* TODO: experiment
blanchet@38986
   268
fun debug_axiom_weight const_tab relevant_consts axiom_consts =
blanchet@38986
   269
  case axiom_consts |> List.partition (pseudoconst_mem I relevant_consts)
blanchet@38986
   270
                    ||> filter_out (pseudoconst_mem swap relevant_consts) of
blanchet@38986
   271
    ([], []) => 0.0
blanchet@38986
   272
  | (_, []) => 1.0
blanchet@38986
   273
  | (rel, irrel) =>
blanchet@38986
   274
    let
blanchet@38986
   275
val _ = tracing (PolyML.makestring ("REL: ", rel))
blanchet@38986
   276
val _ = tracing (PolyML.makestring ("IRREL: ", irrel))
blanchet@38986
   277
      val rel_weight = fold (curry Real.+ o rel_weight const_tab) rel 0.0
blanchet@38986
   278
      val irrel_weight = fold (curry Real.+ o irrel_weight const_tab) irrel 0.0
blanchet@38986
   279
      val res = rel_weight / (rel_weight + irrel_weight)
blanchet@38986
   280
    in if Real.isFinite res then res else 0.0 end
blanchet@38986
   281
*)
blanchet@38986
   282
blanchet@38986
   283
fun pseudoconsts_of_term thy t =
blanchet@38981
   284
  Symtab.fold (fn (x, ys) => fold (fn y => cons (x, y)) ys)
blanchet@38986
   285
              (get_pseudoconsts thy true (SOME true) [t]) []
blanchet@38926
   286
fun pair_consts_axiom theory_relevant thy axiom =
blanchet@37616
   287
  (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
blanchet@38986
   288
                |> pseudoconsts_of_term thy)
paulson@24287
   289
blanchet@38938
   290
type annotated_thm =
blanchet@38983
   291
  ((unit -> string * bool) * thm) * (string * pseudotype list) list
blanchet@37505
   292
blanchet@38986
   293
fun take_most_relevant max_max_imperfect max_relevant remaining_max
blanchet@38986
   294
                       (candidates : (annotated_thm * real) list) =
blanchet@38983
   295
  let
blanchet@38986
   296
    val max_imperfect =
blanchet@38986
   297
      Real.ceil (Math.pow (max_max_imperfect,
blanchet@38986
   298
                           Real.fromInt remaining_max
blanchet@38986
   299
                           / Real.fromInt max_relevant))
blanchet@38986
   300
    val (perfect, imperfect) =
blanchet@38986
   301
      candidates |> List.partition (fn (_, w) => w > 0.99999)
blanchet@38984
   302
                 ||> sort (Real.compare o swap o pairself snd)
blanchet@38986
   303
    val ((accepts, more_rejects), rejects) =
blanchet@38986
   304
      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
blanchet@38983
   305
  in
blanchet@38983
   306
    trace_msg (fn () => "Number of candidates: " ^
blanchet@38984
   307
                        string_of_int (length candidates));
blanchet@38983
   308
    trace_msg (fn () => "Effective threshold: " ^
blanchet@38984
   309
                        Real.toString (#2 (hd accepts)));
blanchet@38986
   310
    trace_msg (fn () => "Actually passed (" ^ Int.toString (length accepts) ^
blanchet@38986
   311
        "): " ^ (accepts
blanchet@38986
   312
                 |> map (fn (((name, _), _), weight) =>
blanchet@38984
   313
                            fst (name ()) ^ " [" ^ Real.toString weight ^ "]")
blanchet@38984
   314
                 |> commas));
blanchet@38986
   315
    (accepts, more_rejects @ rejects)
blanchet@38983
   316
  end
paulson@24287
   317
blanchet@38922
   318
val threshold_divisor = 2.0
blanchet@38922
   319
val ridiculous_threshold = 0.1
blanchet@38986
   320
val max_max_imperfect_fudge_factor = 0.66
blanchet@38922
   321
blanchet@38984
   322
fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
blanchet@38984
   323
                     ({add, del, ...} : relevance_override) axioms goal_ts =
blanchet@38978
   324
  let
blanchet@38978
   325
    val thy = ProofContext.theory_of ctxt
blanchet@38978
   326
    val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
blanchet@38978
   327
                         Symtab.empty
blanchet@38978
   328
    val add_thms = maps (ProofContext.get_fact ctxt) add
blanchet@38978
   329
    val del_thms = maps (ProofContext.get_fact ctxt) del
blanchet@38986
   330
    val max_max_imperfect =
blanchet@38986
   331
      Math.sqrt (Real.fromInt max_relevant * max_max_imperfect_fudge_factor)
blanchet@38986
   332
    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
blanchet@38978
   333
      let
blanchet@38983
   334
        fun game_over rejects =
blanchet@38986
   335
          (* Add "add:" facts. *)
blanchet@38986
   336
          if null add_thms then
blanchet@38986
   337
            []
blanchet@38983
   338
          else
blanchet@38986
   339
            map_filter (fn ((p as (_, th), _), _) =>
blanchet@38986
   340
                           if member Thm.eq_thm add_thms th then SOME p
blanchet@38986
   341
                           else NONE) rejects
blanchet@38986
   342
        fun relevant [] rejects hopeless [] =
blanchet@38986
   343
            (* Nothing has been added this iteration. *)
blanchet@38986
   344
            if j = 0 andalso threshold >= ridiculous_threshold then
blanchet@38986
   345
              (* First iteration? Try again. *)
blanchet@38986
   346
              iter 0 max_relevant (threshold / threshold_divisor) rel_const_tab
blanchet@38986
   347
                   hopeless hopeful
blanchet@38983
   348
            else
blanchet@38986
   349
              game_over (rejects @ hopeless)
blanchet@38986
   350
          | relevant candidates rejects hopeless [] =
blanchet@38978
   351
            let
blanchet@38986
   352
              val (accepts, more_rejects) =
blanchet@38986
   353
                take_most_relevant max_max_imperfect max_relevant remaining_max
blanchet@38986
   354
                                   candidates
blanchet@38978
   355
              val rel_const_tab' =
blanchet@38984
   356
                rel_const_tab
blanchet@38986
   357
                |> fold (add_pseudoconst_to_table false)
blanchet@38986
   358
                        (maps (snd o fst) accepts)
blanchet@38983
   359
              fun is_dirty (c, _) =
blanchet@38983
   360
                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
blanchet@38984
   361
              val (hopeful_rejects, hopeless_rejects) =
blanchet@38984
   362
                 (rejects @ hopeless, ([], []))
blanchet@38984
   363
                 |-> fold (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   364
                              if exists is_dirty consts then
blanchet@38984
   365
                                apfst (cons (ax, NONE))
blanchet@38984
   366
                              else
blanchet@38984
   367
                                apsnd (cons (ax, old_weight)))
blanchet@38984
   368
                 |>> append (more_rejects
blanchet@38984
   369
                             |> map (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   370
                                        (ax, if exists is_dirty consts then NONE
blanchet@38984
   371
                                             else SOME old_weight)))
blanchet@38986
   372
              val threshold =
blanchet@38986
   373
                threshold + (1.0 - threshold)
blanchet@38986
   374
                * Math.pow (decay, Real.fromInt (length accepts))
blanchet@38986
   375
              val remaining_max = remaining_max - length accepts
blanchet@38978
   376
            in
blanchet@38983
   377
              trace_msg (fn () => "New or updated constants: " ^
blanchet@38983
   378
                  commas (rel_const_tab' |> Symtab.dest
blanchet@38983
   379
                          |> subtract (op =) (Symtab.dest rel_const_tab)
blanchet@38983
   380
                          |> map string_for_super_pseudoconst));
blanchet@38984
   381
              map (fst o fst) accepts @
blanchet@38986
   382
              (if remaining_max = 0 then
blanchet@38984
   383
                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
blanchet@38984
   384
               else
blanchet@38986
   385
                 iter (j + 1) remaining_max threshold rel_const_tab'
blanchet@38986
   386
                      hopeless_rejects hopeful_rejects)
blanchet@38978
   387
            end
blanchet@38986
   388
          | relevant candidates rejects hopeless
blanchet@38978
   389
                     (((ax as ((name, th), axiom_consts)), cached_weight)
blanchet@38986
   390
                      :: hopeful) =
blanchet@38978
   391
            let
blanchet@38978
   392
              val weight =
blanchet@38978
   393
                case cached_weight of
blanchet@38978
   394
                  SOME w => w
blanchet@38978
   395
                | NONE => axiom_weight const_tab rel_const_tab axiom_consts
blanchet@38986
   396
(* TODO: experiment
blanchet@38986
   397
val _ = if String.isPrefix "lift.simps(3" (fst (name ())) then
blanchet@38986
   398
tracing ("*** " ^ (fst (name ())) ^ PolyML.makestring (debug_axiom_weight const_tab rel_const_tab axiom_consts))
blanchet@38986
   399
else
blanchet@38986
   400
()
blanchet@38986
   401
*)
blanchet@38978
   402
            in
blanchet@38980
   403
              if weight >= threshold then
blanchet@38986
   404
                relevant ((ax, weight) :: candidates) rejects hopeless hopeful
blanchet@38978
   405
              else
blanchet@38986
   406
                relevant candidates ((ax, weight) :: rejects) hopeless hopeful
blanchet@38978
   407
            end
blanchet@38978
   408
        in
blanchet@38983
   409
          trace_msg (fn () =>
blanchet@38983
   410
              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
blanchet@38983
   411
              Real.toString threshold ^ ", constants: " ^
blanchet@38983
   412
              commas (rel_const_tab |> Symtab.dest
blanchet@38983
   413
                      |> filter (curry (op <>) [] o snd)
blanchet@38983
   414
                      |> map string_for_super_pseudoconst));
blanchet@38986
   415
          relevant [] [] hopeless hopeful
blanchet@38978
   416
        end
blanchet@38978
   417
  in
blanchet@38978
   418
    axioms |> filter_out (member Thm.eq_thm del_thms o snd)
blanchet@38978
   419
           |> map (rpair NONE o pair_consts_axiom theory_relevant thy)
blanchet@38984
   420
           |> iter 0 max_relevant threshold0
blanchet@38986
   421
                   (get_pseudoconsts thy false (SOME false) goal_ts) []
blanchet@38978
   422
           |> tap (fn res => trace_msg (fn () =>
blanchet@38925
   423
                                "Total relevant: " ^ Int.toString (length res)))
blanchet@38978
   424
  end
paulson@24287
   425
blanchet@38983
   426
paulson@24287
   427
(***************************************************************)
mengj@19768
   428
(* Retrieving and filtering lemmas                             *)
mengj@19768
   429
(***************************************************************)
mengj@19768
   430
paulson@33022
   431
(*** retrieve lemmas and filter them ***)
mengj@19768
   432
paulson@20757
   433
(*Reject theorems with names like "List.filter.filter_list_def" or
paulson@21690
   434
  "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
paulson@20757
   435
fun is_package_def a =
wenzelm@30364
   436
  let val names = Long_Name.explode a
paulson@21690
   437
  in
paulson@21690
   438
     length names > 2 andalso
paulson@21690
   439
     not (hd names = "local") andalso
paulson@21690
   440
     String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
paulson@21690
   441
  end;
paulson@20757
   442
blanchet@38331
   443
fun make_fact_table xs =
blanchet@37616
   444
  fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
blanchet@38331
   445
fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
mengj@19768
   446
blanchet@37626
   447
(* FIXME: put other record thms here, or declare as "no_atp" *)
blanchet@37626
   448
val multi_base_blacklist =
blanchet@37626
   449
  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
blanchet@38921
   450
   "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
blanchet@38921
   451
   "case_cong", "weak_case_cong"]
blanchet@38921
   452
  |> map (prefix ".")
blanchet@37626
   453
blanchet@37626
   454
val max_lambda_nesting = 3
blanchet@37626
   455
blanchet@37626
   456
fun term_has_too_many_lambdas max (t1 $ t2) =
blanchet@37626
   457
    exists (term_has_too_many_lambdas max) [t1, t2]
blanchet@37626
   458
  | term_has_too_many_lambdas max (Abs (_, _, t)) =
blanchet@37626
   459
    max = 0 orelse term_has_too_many_lambdas (max - 1) t
blanchet@37626
   460
  | term_has_too_many_lambdas _ _ = false
blanchet@37626
   461
blanchet@37626
   462
(* Don't count nested lambdas at the level of formulas, since they are
blanchet@37626
   463
   quantifiers. *)
blanchet@37626
   464
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
blanchet@37626
   465
    formula_has_too_many_lambdas (T :: Ts) t
blanchet@37626
   466
  | formula_has_too_many_lambdas Ts t =
blanchet@37626
   467
    if is_formula_type (fastype_of1 (Ts, t)) then
blanchet@37626
   468
      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
blanchet@37626
   469
    else
blanchet@37626
   470
      term_has_too_many_lambdas max_lambda_nesting t
blanchet@37626
   471
blanchet@38931
   472
(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
blanchet@37626
   473
   was 11. *)
blanchet@37626
   474
val max_apply_depth = 15
blanchet@37626
   475
blanchet@37626
   476
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
blanchet@37626
   477
  | apply_depth (Abs (_, _, t)) = apply_depth t
blanchet@37626
   478
  | apply_depth _ = 0
blanchet@37626
   479
blanchet@37626
   480
fun is_formula_too_complex t =
blanchet@38331
   481
  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
blanchet@37626
   482
blanchet@37539
   483
val exists_sledgehammer_const =
blanchet@37626
   484
  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
blanchet@37626
   485
blanchet@38890
   486
fun is_strange_theorem th =
blanchet@37626
   487
  case head_of (concl_of th) of
blanchet@37626
   488
      Const (a, _) => (a <> @{const_name Trueprop} andalso
blanchet@37626
   489
                       a <> @{const_name "=="})
blanchet@37626
   490
    | _ => false
blanchet@37626
   491
blanchet@37626
   492
val type_has_top_sort =
blanchet@37626
   493
  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
blanchet@37626
   494
blanchet@38331
   495
(**** Predicates to detect unwanted facts (prolific or likely to cause
blanchet@37347
   496
      unsoundness) ****)
paulson@21470
   497
blanchet@38513
   498
(* Too general means, positive equality literal with a variable X as one
blanchet@38513
   499
   operand, when X does not occur properly in the other operand. This rules out
blanchet@38513
   500
   clearly inconsistent facts such as X = a | X = b, though it by no means
blanchet@38513
   501
   guarantees soundness. *)
paulson@21470
   502
blanchet@38513
   503
(* Unwanted equalities are those between a (bound or schematic) variable that
blanchet@38513
   504
   does not properly occur in the second operand. *)
blanchet@38830
   505
val is_exhaustive_finite =
blanchet@38830
   506
  let
blanchet@38852
   507
    fun is_bad_equal (Var z) t =
blanchet@38852
   508
        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
blanchet@38852
   509
      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
blanchet@38852
   510
      | is_bad_equal _ _ = false
blanchet@38852
   511
    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
blanchet@38830
   512
    fun do_formula pos t =
blanchet@38830
   513
      case (pos, t) of
blanchet@38838
   514
        (_, @{const Trueprop} $ t1) => do_formula pos t1
blanchet@38830
   515
      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
blanchet@38830
   516
        do_formula pos t'
blanchet@38830
   517
      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
blanchet@38830
   518
        do_formula pos t'
blanchet@38830
   519
      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
blanchet@38830
   520
        do_formula pos t'
blanchet@38830
   521
      | (_, @{const "==>"} $ t1 $ t2) =>
blanchet@38852
   522
        do_formula (not pos) t1 andalso
blanchet@38852
   523
        (t2 = @{prop False} orelse do_formula pos t2)
blanchet@38830
   524
      | (_, @{const "op -->"} $ t1 $ t2) =>
blanchet@38852
   525
        do_formula (not pos) t1 andalso
blanchet@38852
   526
        (t2 = @{const False} orelse do_formula pos t2)
blanchet@38830
   527
      | (_, @{const Not} $ t1) => do_formula (not pos) t1
blanchet@38830
   528
      | (true, @{const "op |"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@38830
   529
      | (false, @{const "op &"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@38830
   530
      | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@38830
   531
      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@38830
   532
      | _ => false
blanchet@38830
   533
  in do_formula true end
blanchet@38830
   534
blanchet@38815
   535
fun has_bound_or_var_of_type tycons =
blanchet@38815
   536
  exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
blanchet@38815
   537
                   | Abs (_, Type (s, _), _) => member (op =) tycons s
blanchet@38815
   538
                   | _ => false)
paulson@21431
   539
blanchet@38331
   540
(* Facts are forbidden to contain variables of these types. The typical reason
blanchet@37347
   541
   is that they lead to unsoundness. Note that "unit" satisfies numerous
blanchet@38331
   542
   equations like "?x = ()". The resulting clauses will have no type constraint,
blanchet@37347
   543
   yielding false proofs. Even "bool" leads to many unsound proofs, though only
blanchet@37347
   544
   for higher-order problems. *)
blanchet@38815
   545
val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
paulson@22217
   546
blanchet@38331
   547
(* Facts containing variables of type "unit" or "bool" or of the form
blanchet@38514
   548
   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
blanchet@38514
   549
   are omitted. *)
blanchet@38816
   550
fun is_dangerous_term full_types t =
blanchet@38832
   551
  not full_types andalso
blanchet@38918
   552
  let val t = transform_elim_term t in
blanchet@38918
   553
    has_bound_or_var_of_type dangerous_types t orelse
blanchet@38918
   554
    is_exhaustive_finite t
blanchet@38918
   555
  end
paulson@21470
   556
blanchet@38850
   557
fun is_theorem_bad_for_atps full_types thm =
blanchet@38850
   558
  let val t = prop_of thm in
blanchet@38850
   559
    is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
blanchet@38850
   560
    is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
blanchet@38890
   561
    is_strange_theorem thm
blanchet@38850
   562
  end
blanchet@38850
   563
blanchet@38935
   564
fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
blanchet@38850
   565
  let
blanchet@38936
   566
    val is_chained = member Thm.eq_thm chained_ths
blanchet@38936
   567
    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt)
blanchet@38882
   568
    val local_facts = ProofContext.facts_of ctxt
blanchet@38882
   569
    val named_locals = local_facts |> Facts.dest_static []
blanchet@38936
   570
    (* Unnamed, not chained formulas with schematic variables are omitted,
blanchet@38936
   571
       because they are rejected by the backticks (`...`) parser for some
blanchet@38936
   572
       reason. *)
blanchet@38977
   573
    fun is_good_unnamed_local th =
blanchet@38977
   574
      forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
blanchet@38977
   575
      andalso (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th))
blanchet@38882
   576
    val unnamed_locals =
blanchet@38977
   577
      local_facts |> Facts.props |> filter is_good_unnamed_local
blanchet@38936
   578
                  |> map (pair "" o single)
blanchet@38850
   579
    val full_space =
blanchet@38977
   580
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
blanchet@38936
   581
    fun add_valid_facts foldx facts =
blanchet@38938
   582
      foldx (fn (name0, ths) =>
blanchet@38938
   583
        if name0 <> "" andalso
blanchet@38938
   584
           forall (not o member Thm.eq_thm add_thms) ths andalso
blanchet@38938
   585
           (Facts.is_concealed facts name0 orelse
blanchet@38938
   586
            (respect_no_atp andalso is_package_def name0) orelse
blanchet@38938
   587
            exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
blanchet@38938
   588
            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
blanchet@38850
   589
          I
blanchet@38850
   590
        else
blanchet@38850
   591
          let
blanchet@38938
   592
            val multi = length ths > 1
blanchet@38935
   593
            fun backquotify th =
blanchet@38935
   594
              "`" ^ Print_Mode.setmp [Print_Mode.input]
blanchet@38935
   595
                                 (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
blanchet@38977
   596
              |> String.translate (fn c => if Char.isPrint c then str c else "")
blanchet@38977
   597
              |> simplify_spaces
blanchet@38938
   598
            fun check_thms a =
blanchet@38938
   599
              case try (ProofContext.get_thms ctxt) a of
blanchet@38938
   600
                NONE => false
blanchet@38938
   601
              | SOME ths' => Thm.eq_thms (ths, ths')
blanchet@38850
   602
          in
blanchet@38938
   603
            pair 1
blanchet@38938
   604
            #> fold (fn th => fn (j, rest) =>
blanchet@38938
   605
                 (j + 1,
blanchet@38938
   606
                  if is_theorem_bad_for_atps full_types th andalso
blanchet@38938
   607
                     not (member Thm.eq_thm add_thms th) then
blanchet@38938
   608
                    rest
blanchet@38938
   609
                  else
blanchet@38938
   610
                    (fn () =>
blanchet@38938
   611
                        (if name0 = "" then
blanchet@38938
   612
                           th |> backquotify
blanchet@38938
   613
                         else
blanchet@38938
   614
                           let
blanchet@38938
   615
                             val name1 = Facts.extern facts name0
blanchet@38938
   616
                             val name2 = Name_Space.extern full_space name0
blanchet@38938
   617
                           in
blanchet@38938
   618
                             case find_first check_thms [name1, name2, name0] of
blanchet@38983
   619
                               SOME name => repair_name reserved multi j name
blanchet@38938
   620
                             | NONE => ""
blanchet@38938
   621
                           end, is_chained th), (multi, th)) :: rest)) ths
blanchet@38938
   622
            #> snd
blanchet@38850
   623
          end)
blanchet@38882
   624
  in
blanchet@38927
   625
    [] |> add_valid_facts fold local_facts (unnamed_locals @ named_locals)
blanchet@38927
   626
       |> add_valid_facts Facts.fold_static global_facts global_facts
blanchet@38882
   627
  end
blanchet@38850
   628
blanchet@38850
   629
(* The single-name theorems go after the multiple-name ones, so that single
blanchet@38850
   630
   names are preferred when both are available. *)
blanchet@38938
   631
fun name_thm_pairs ctxt respect_no_atp =
blanchet@38983
   632
  List.partition (fst o snd) #> op @ #> map (apsnd snd)
blanchet@38938
   633
  #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
blanchet@38850
   634
blanchet@38850
   635
(***************************************************************)
blanchet@38850
   636
(* ATP invocation methods setup                                *)
blanchet@38850
   637
(***************************************************************)
blanchet@38850
   638
blanchet@38984
   639
fun relevant_facts full_types (threshold0, threshold1) max_relevant
blanchet@38983
   640
                   theory_relevant (relevance_override as {add, del, only})
blanchet@38229
   641
                   (ctxt, (chained_ths, _)) hyp_ts concl_t =
blanchet@37534
   642
  let
blanchet@38984
   643
    val decay = 1.0 - Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
blanchet@38984
   644
                                1.0 / Real.fromInt (max_relevant + 1))
blanchet@37534
   645
    val add_thms = maps (ProofContext.get_fact ctxt) add
blanchet@38935
   646
    val reserved = reserved_isar_keyword_table ()
blanchet@37534
   647
    val axioms =
blanchet@38938
   648
      (if only then
blanchet@38983
   649
         maps (name_thm_pairs_from_ref ctxt reserved chained_ths) add
blanchet@38938
   650
       else
blanchet@38938
   651
         all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
blanchet@38927
   652
      |> name_thm_pairs ctxt (respect_no_atp andalso not only)
blanchet@38818
   653
      |> make_unique
blanchet@37534
   654
  in
blanchet@38927
   655
    trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
blanchet@38927
   656
                        " theorems");
blanchet@38984
   657
    (if threshold0 > 1.0 orelse threshold0 > threshold1 then
blanchet@38978
   658
       []
blanchet@38984
   659
     else if threshold0 < 0.0 then
blanchet@38978
   660
       axioms
blanchet@38978
   661
     else
blanchet@38984
   662
       relevance_filter ctxt threshold0 decay max_relevant theory_relevant
blanchet@38984
   663
                        relevance_override axioms (concl_t :: hyp_ts))
blanchet@38983
   664
    |> map (apfst (fn f => f ())) |> sort_wrt (fst o fst)
blanchet@37534
   665
  end
immler@30536
   666
paulson@15347
   667
end;