src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML
author wenzelm
Thu, 09 Jun 2011 16:34:49 +0200
changeset 44206 2b47822868e4
parent 44086 cef69d31bfa4
child 44217 b19d95b4d736
permissions -rw-r--r--
discontinued Name.variant to emphasize that this is old-style / indirect;
blanchet@39232
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter.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@40139
     5
Sledgehammer's relevance filter.
wenzelm@33317
     6
*)
paulson@15452
     7
blanchet@39232
     8
signature SLEDGEHAMMER_FILTER =
wenzelm@16802
     9
sig
blanchet@43926
    10
  type locality = ATP_Translate.locality
blanchet@38991
    11
blanchet@40251
    12
  type relevance_fudge =
blanchet@43603
    13
    {local_const_multiplier : real,
blanchet@41407
    14
     worse_irrel_freq : real,
blanchet@40251
    15
     higher_order_irrel_weight : real,
blanchet@40251
    16
     abs_rel_weight : real,
blanchet@40251
    17
     abs_irrel_weight : real,
blanchet@40251
    18
     skolem_irrel_weight : real,
blanchet@40251
    19
     theory_const_rel_weight : real,
blanchet@40251
    20
     theory_const_irrel_weight : real,
blanchet@43600
    21
     chained_const_irrel_weight : real,
blanchet@40251
    22
     intro_bonus : real,
blanchet@40251
    23
     elim_bonus : real,
blanchet@40251
    24
     simp_bonus : real,
blanchet@40251
    25
     local_bonus : real,
blanchet@40251
    26
     assum_bonus : real,
blanchet@40251
    27
     chained_bonus : real,
blanchet@40251
    28
     max_imperfect : real,
blanchet@40251
    29
     max_imperfect_exp : real,
blanchet@40251
    30
     threshold_divisor : real,
blanchet@40251
    31
     ridiculous_threshold : real}
blanchet@40251
    32
blanchet@35964
    33
  type relevance_override =
blanchet@40251
    34
    {add : (Facts.ref * Attrib.src list) list,
blanchet@40251
    35
     del : (Facts.ref * Attrib.src list) list,
blanchet@40251
    36
     only : bool}
blanchet@35964
    37
blanchet@43517
    38
  val trace : bool Config.T
blanchet@43593
    39
  val new_monomorphizer : bool Config.T
blanchet@43593
    40
  val ignore_no_atp : bool Config.T
blanchet@43593
    41
  val instantiate_inducts : bool Config.T
blanchet@40446
    42
  val fact_from_ref :
blanchet@39240
    43
    Proof.context -> unit Symtab.table -> thm list
blanchet@39240
    44
    -> Facts.ref * Attrib.src list -> ((string * locality) * thm) list
blanchet@42638
    45
  val all_facts :
blanchet@43785
    46
    Proof.context -> 'a Symtab.table -> bool -> (term -> bool) -> thm list
blanchet@44086
    47
    -> thm list -> (((unit -> string) * locality) * thm) list
blanchet@42639
    48
  val const_names_in_fact :
blanchet@42639
    49
    theory -> (string * typ -> term list -> bool * term list) -> term
blanchet@42639
    50
    -> string list
blanchet@37347
    51
  val relevant_facts :
blanchet@43785
    52
    Proof.context -> real * real -> int -> (term -> bool)
blanchet@41584
    53
    -> (string * typ -> term list -> bool * term list) -> relevance_fudge
blanchet@41314
    54
    -> relevance_override -> thm list -> term list -> term
blanchet@40251
    55
    -> ((string * locality) * thm) list
paulson@15347
    56
end;
paulson@15347
    57
blanchet@39232
    58
structure Sledgehammer_Filter : SLEDGEHAMMER_FILTER =
paulson@15347
    59
struct
paulson@15347
    60
blanchet@43926
    61
open ATP_Translate
blanchet@38890
    62
open Sledgehammer_Util
blanchet@38890
    63
blanchet@43517
    64
val trace =
blanchet@43517
    65
  Attrib.setup_config_bool @{binding sledgehammer_filter_trace} (K false)
blanchet@43517
    66
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@35826
    67
blanchet@41521
    68
(* experimental features *)
blanchet@43593
    69
val new_monomorphizer =
blanchet@43593
    70
  Attrib.setup_config_bool @{binding sledgehammer_new_monomorphizer} (K false)
blanchet@43593
    71
val ignore_no_atp =
blanchet@43593
    72
  Attrib.setup_config_bool @{binding sledgehammer_ignore_no_atp} (K false)
blanchet@43593
    73
val instantiate_inducts =
blanchet@43593
    74
  Attrib.setup_config_bool @{binding sledgehammer_instantiate_inducts} (K false)
blanchet@39062
    75
blanchet@40251
    76
type relevance_fudge =
blanchet@43603
    77
  {local_const_multiplier : real,
blanchet@41407
    78
   worse_irrel_freq : real,
blanchet@40251
    79
   higher_order_irrel_weight : real,
blanchet@40251
    80
   abs_rel_weight : real,
blanchet@40251
    81
   abs_irrel_weight : real,
blanchet@40251
    82
   skolem_irrel_weight : real,
blanchet@40251
    83
   theory_const_rel_weight : real,
blanchet@40251
    84
   theory_const_irrel_weight : real,
blanchet@43600
    85
   chained_const_irrel_weight : real,
blanchet@40251
    86
   intro_bonus : real,
blanchet@40251
    87
   elim_bonus : real,
blanchet@40251
    88
   simp_bonus : real,
blanchet@40251
    89
   local_bonus : real,
blanchet@40251
    90
   assum_bonus : real,
blanchet@40251
    91
   chained_bonus : real,
blanchet@40251
    92
   max_imperfect : real,
blanchet@40251
    93
   max_imperfect_exp : real,
blanchet@40251
    94
   threshold_divisor : real,
blanchet@40251
    95
   ridiculous_threshold : real}
blanchet@40251
    96
blanchet@35964
    97
type relevance_override =
blanchet@40251
    98
  {add : (Facts.ref * Attrib.src list) list,
blanchet@40251
    99
   del : (Facts.ref * Attrib.src list) list,
blanchet@40251
   100
   only : bool}
paulson@21070
   101
blanchet@37616
   102
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
blanchet@40077
   103
val abs_name = sledgehammer_prefix ^ "abs"
blanchet@40077
   104
val skolem_prefix = sledgehammer_prefix ^ "sko"
blanchet@39236
   105
val theory_const_suffix = Long_Name.separator ^ " 1"
blanchet@37616
   106
blanchet@40468
   107
fun needs_quoting reserved s =
blanchet@40621
   108
  Symtab.defined reserved s orelse
wenzelm@43162
   109
  exists (not o Lexicon.is_identifier) (Long_Name.explode s)
blanchet@40468
   110
blanchet@40621
   111
fun make_name reserved multi j name =
blanchet@40468
   112
  (name |> needs_quoting reserved name ? quote) ^
wenzelm@41739
   113
  (if multi then "(" ^ string_of_int j ^ ")" else "")
blanchet@38983
   114
blanchet@40621
   115
fun explode_interval _ (Facts.FromTo (i, j)) = i upto j
blanchet@40621
   116
  | explode_interval max (Facts.From i) = i upto i + max - 1
blanchet@40621
   117
  | explode_interval _ (Facts.Single i) = [i]
blanchet@40621
   118
blanchet@41527
   119
val backquote =
blanchet@41527
   120
  raw_explode #> map (fn "`" => "\\`" | s => s) #> implode #> enclose "`" "`"
blanchet@40446
   121
fun fact_from_ref ctxt reserved chained_ths (xthm as (xref, args)) =
blanchet@38983
   122
  let
blanchet@39240
   123
    val ths = Attrib.eval_thms ctxt [xthm]
blanchet@39240
   124
    val bracket =
blanchet@42870
   125
      map (enclose "[" "]" o Pretty.str_of o Args.pretty_src ctxt) args
blanchet@42870
   126
      |> implode
blanchet@40621
   127
    fun nth_name j =
blanchet@39240
   128
      case xref of
blanchet@41527
   129
        Facts.Fact s => backquote s ^ bracket
blanchet@39240
   130
      | Facts.Named (("", _), _) => "[" ^ bracket ^ "]"
blanchet@40621
   131
      | Facts.Named ((name, _), NONE) =>
blanchet@40621
   132
        make_name reserved (length ths > 1) (j + 1) name ^ bracket
blanchet@40621
   133
      | Facts.Named ((name, _), SOME intervals) =>
blanchet@40621
   134
        make_name reserved true
blanchet@40621
   135
                 (nth (maps (explode_interval (length ths)) intervals) j) name ^
blanchet@40621
   136
        bracket
blanchet@38983
   137
  in
blanchet@40621
   138
    (ths, (0, []))
blanchet@38991
   139
    |-> fold (fn th => fn (j, rest) =>
blanchet@40621
   140
                 (j + 1, ((nth_name j,
blanchet@43798
   141
                          if member Thm.eq_thm_prop chained_ths th then Chained
blanchet@38991
   142
                          else General), th) :: rest))
blanchet@38983
   143
    |> snd
blanchet@38938
   144
  end
blanchet@37616
   145
blanchet@41447
   146
(* This is a terrible hack. Free variables are sometimes code as "M__" when they
blanchet@41447
   147
   are displayed as "M" and we want to avoid clashes with these. But sometimes
blanchet@41447
   148
   it's even worse: "Ma__" encodes "M". So we simply reserve all prefixes of all
blanchet@41447
   149
   free variables. In the worse case scenario, where the fact won't be resolved
blanchet@41447
   150
   correctly, the user can fix it manually, e.g., by naming the fact in
blanchet@41447
   151
   question. Ideally we would need nothing of it, but backticks just don't work
blanchet@41447
   152
   with schematic variables. *)
blanchet@41447
   153
fun all_prefixes_of s =
blanchet@41447
   154
  map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1)
blanchet@41447
   155
fun close_form t =
blanchet@41447
   156
  (t, [] |> Term.add_free_names t |> maps all_prefixes_of)
blanchet@41447
   157
  |> fold (fn ((s, i), T) => fn (t', taken) =>
wenzelm@44206
   158
              let val s' = singleton (Name.variant_list taken) s in
blanchet@41447
   159
                ((if fastype_of t' = HOLogic.boolT then HOLogic.all_const
blanchet@41447
   160
                  else Term.all) T
blanchet@41447
   161
                 $ Abs (s', T, abstract_over (Var ((s, i), T), t')),
blanchet@41447
   162
                 s' :: taken)
blanchet@41447
   163
              end)
blanchet@41447
   164
          (Term.add_vars t [] |> sort_wrt (fst o fst))
blanchet@41447
   165
  |> fst
blanchet@41447
   166
blanchet@41447
   167
fun string_for_term ctxt t =
blanchet@41447
   168
  Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
blanchet@41447
   169
                   (print_mode_value ())) (Syntax.string_of_term ctxt) t
blanchet@41447
   170
  |> String.translate (fn c => if Char.isPrint c then str c else "")
blanchet@41447
   171
  |> simplify_spaces
blanchet@41447
   172
blanchet@41447
   173
(** Structural induction rules **)
blanchet@41447
   174
blanchet@41448
   175
fun struct_induct_rule_on th =
blanchet@41447
   176
  case Logic.strip_horn (prop_of th) of
blanchet@41447
   177
    (prems, @{const Trueprop}
blanchet@41447
   178
            $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
blanchet@41447
   179
    if not (is_TVar ind_T) andalso length prems > 1 andalso
blanchet@41447
   180
       exists (exists_subterm (curry (op aconv) p)) prems andalso
blanchet@41447
   181
       not (exists (exists_subterm (curry (op aconv) a)) prems) then
blanchet@41447
   182
      SOME (p_name, ind_T)
blanchet@41447
   183
    else
blanchet@41447
   184
      NONE
blanchet@41447
   185
  | _ => NONE
blanchet@41447
   186
blanchet@44086
   187
fun instantiate_induct_rule ctxt concl_prop p_name ((name, loc), th) ind_x =
blanchet@41447
   188
  let
blanchet@41447
   189
    fun varify_noninducts (t as Free (s, T)) =
blanchet@41447
   190
        if (s, T) = ind_x orelse can dest_funT T then t else Var ((s, 0), T)
blanchet@41447
   191
      | varify_noninducts t = t
blanchet@41447
   192
    val p_inst =
blanchet@41447
   193
      concl_prop |> map_aterms varify_noninducts |> close_form
blanchet@41447
   194
                 |> lambda (Free ind_x)
blanchet@41447
   195
                 |> string_for_term ctxt
blanchet@41447
   196
  in
blanchet@41455
   197
    ((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]", loc),
blanchet@44086
   198
     th |> read_instantiate ctxt [((p_name, 0), p_inst)])
blanchet@41447
   199
  end
blanchet@41447
   200
blanchet@41447
   201
fun type_match thy (T1, T2) =
blanchet@41447
   202
  (Sign.typ_match thy (T2, T1) Vartab.empty; true)
blanchet@41447
   203
  handle Type.TYPE_MATCH => false
blanchet@41447
   204
blanchet@44086
   205
fun instantiate_if_induct_rule ctxt stmt stmt_xs (ax as (_, th)) =
blanchet@41448
   206
  case struct_induct_rule_on th of
blanchet@41447
   207
    SOME (p_name, ind_T) =>
wenzelm@43232
   208
    let val thy = Proof_Context.theory_of ctxt in
blanchet@41447
   209
      stmt_xs |> filter (fn (_, T) => type_match thy (T, ind_T))
blanchet@41447
   210
              |> map_filter (try (instantiate_induct_rule ctxt stmt p_name ax))
blanchet@41447
   211
    end
blanchet@41447
   212
  | NONE => [ax]
blanchet@41447
   213
wenzelm@28477
   214
(***************************************************************)
wenzelm@28477
   215
(* Relevance Filtering                                         *)
wenzelm@28477
   216
(***************************************************************)
mengj@19194
   217
paulson@24287
   218
(*** constants with types ***)
paulson@24287
   219
blanchet@39165
   220
fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) =
blanchet@39165
   221
    order_of_type T1 (* cheat: pretend sets are first-order *)
blanchet@39165
   222
  | order_of_type (Type (@{type_name fun}, [T1, T2])) =
blanchet@39165
   223
    Int.max (order_of_type T1 + 1, order_of_type T2)
blanchet@39165
   224
  | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
blanchet@39165
   225
  | order_of_type _ = 0
blanchet@39165
   226
blanchet@39058
   227
(* An abstraction of Isabelle types and first-order terms *)
blanchet@39058
   228
datatype pattern = PVar | PApp of string * pattern list
blanchet@39165
   229
datatype ptype = PType of int * pattern list
blanchet@38983
   230
blanchet@39058
   231
fun string_for_pattern PVar = "_"
blanchet@39058
   232
  | string_for_pattern (PApp (s, ps)) =
blanchet@39058
   233
    if null ps then s else s ^ string_for_patterns ps
blanchet@39058
   234
and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
blanchet@39165
   235
fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
paulson@24287
   236
paulson@24287
   237
(*Is the second type an instance of the first one?*)
blanchet@39059
   238
fun match_pattern (PVar, _) = true
blanchet@39059
   239
  | match_pattern (PApp _, PVar) = false
blanchet@39059
   240
  | match_pattern (PApp (s, ps), PApp (t, qs)) =
blanchet@39059
   241
    s = t andalso match_patterns (ps, qs)
blanchet@39059
   242
and match_patterns (_, []) = true
blanchet@39059
   243
  | match_patterns ([], _) = false
blanchet@39059
   244
  | match_patterns (p :: ps, q :: qs) =
blanchet@39059
   245
    match_pattern (p, q) andalso match_patterns (ps, qs)
blanchet@39165
   246
fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
paulson@24287
   247
blanchet@39058
   248
(* Is there a unifiable constant? *)
blanchet@39062
   249
fun pconst_mem f consts (s, ps) =
blanchet@39165
   250
  exists (curry (match_ptype o f) ps)
blanchet@39062
   251
         (map snd (filter (curry (op =) s o fst) consts))
blanchet@39062
   252
fun pconst_hyper_mem f const_tab (s, ps) =
blanchet@39165
   253
  exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
blanchet@37505
   254
blanchet@39165
   255
fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
blanchet@39165
   256
  | pattern_for_type (TFree (s, _)) = PApp (s, [])
blanchet@39165
   257
  | pattern_for_type (TVar _) = PVar
blanchet@39062
   258
blanchet@38983
   259
(* Pairs a constant with the list of its type instantiations. *)
blanchet@41452
   260
fun ptype thy const x =
blanchet@39165
   261
  (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
blanchet@41452
   262
   else [])
blanchet@41452
   263
fun rich_ptype thy const (s, T) =
blanchet@41452
   264
  PType (order_of_type T, ptype thy const (s, T))
blanchet@41452
   265
fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T))
paulson@24287
   266
blanchet@39165
   267
fun string_for_hyper_pconst (s, ps) =
blanchet@39165
   268
  s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
paulson@24287
   269
blanchet@39058
   270
(* Add a pconstant to the table, but a [] entry means a standard
blanchet@39054
   271
   connective, which we ignore.*)
blanchet@41314
   272
fun add_pconst_to_table also_skolem (s, p) =
blanchet@41314
   273
  if (not also_skolem andalso String.isPrefix skolem_prefix s) then I
blanchet@41314
   274
  else Symtab.map_default (s, [p]) (insert (op =) p)
blanchet@39054
   275
blanchet@43606
   276
(* Set constants tend to pull in too many irrelevant facts. We limit the damage
blanchet@43606
   277
   by treating them more or less as if they were built-in but add their
blanchet@43606
   278
   definition at the end. *)
blanchet@43606
   279
val set_consts =
blanchet@43606
   280
  [(@{const_name Collect}, @{thms Collect_def}),
blanchet@43606
   281
   (@{const_name Set.member}, @{thms mem_def})]
blanchet@43606
   282
blanchet@43606
   283
val is_set_const = AList.defined (op =) set_consts
blanchet@43606
   284
blanchet@43597
   285
fun add_pconsts_in_term thy is_built_in_const also_skolems pos =
blanchet@37505
   286
  let
blanchet@39054
   287
    val flip = Option.map not
blanchet@38810
   288
    (* We include free variables, as well as constants, to handle locales. For
blanchet@41453
   289
       each quantifiers that must necessarily be skolemized by the automatic
blanchet@41453
   290
       prover, we introduce a fresh constant to simulate the effect of
blanchet@41453
   291
       Skolemization. *)
blanchet@43606
   292
    fun do_const const ext_arg (x as (s, _)) ts =
blanchet@41584
   293
      let val (built_in, ts) = is_built_in_const x ts in
blanchet@43606
   294
        if is_set_const s then
blanchet@43606
   295
          fold (do_term ext_arg) ts
blanchet@43606
   296
        else
blanchet@43606
   297
          (not built_in
blanchet@43606
   298
           ? add_pconst_to_table also_skolems (rich_pconst thy const x))
blanchet@43606
   299
          #> fold (do_term false) ts
blanchet@41584
   300
      end
blanchet@43606
   301
    and do_term ext_arg t =
blanchet@39062
   302
      case strip_comb t of
blanchet@43606
   303
        (Const x, ts) => do_const true ext_arg x ts
blanchet@43606
   304
      | (Free x, ts) => do_const false ext_arg x ts
blanchet@39165
   305
      | (Abs (_, T, t'), ts) =>
blanchet@43606
   306
        ((null ts andalso not ext_arg)
blanchet@43606
   307
         (* Since lambdas on the right-hand side of equalities are usually
blanchet@43606
   308
            extensionalized later by "extensionalize_term", we don't penalize
blanchet@43606
   309
            them here. *)
blanchet@41314
   310
         ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, [])))
blanchet@43600
   311
        #> fold (do_term false) (t' :: ts)
blanchet@43600
   312
      | (_, ts) => fold (do_term false) ts
blanchet@39165
   313
    fun do_quantifier will_surely_be_skolemized abs_T body_t =
blanchet@37533
   314
      do_formula pos body_t
blanchet@38986
   315
      #> (if also_skolems andalso will_surely_be_skolemized then
blanchet@41314
   316
            add_pconst_to_table true
blanchet@41455
   317
                (gensym skolem_prefix, PType (order_of_type abs_T, []))
blanchet@38810
   318
          else
blanchet@38810
   319
            I)
blanchet@43606
   320
    and do_term_or_formula ext_arg T =
blanchet@43606
   321
      if T = HOLogic.boolT then do_formula NONE else do_term ext_arg
blanchet@37533
   322
    and do_formula pos t =
blanchet@37533
   323
      case t of
blanchet@39165
   324
        Const (@{const_name all}, _) $ Abs (_, T, t') =>
blanchet@39165
   325
        do_quantifier (pos = SOME false) T t'
blanchet@37533
   326
      | @{const "==>"} $ t1 $ t2 =>
blanchet@37533
   327
        do_formula (flip pos) t1 #> do_formula pos t2
blanchet@37533
   328
      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
blanchet@43606
   329
        do_term_or_formula false T t1 #> do_term_or_formula true T t2
blanchet@37533
   330
      | @{const Trueprop} $ t1 => do_formula pos t1
blanchet@41388
   331
      | @{const False} => I
blanchet@41388
   332
      | @{const True} => I
blanchet@37533
   333
      | @{const Not} $ t1 => do_formula (flip pos) t1
blanchet@39165
   334
      | Const (@{const_name All}, _) $ Abs (_, T, t') =>
blanchet@39165
   335
        do_quantifier (pos = SOME false) T t'
blanchet@39165
   336
      | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
blanchet@39165
   337
        do_quantifier (pos = SOME true) T t'
haftmann@39028
   338
      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
haftmann@39028
   339
      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
haftmann@39019
   340
      | @{const HOL.implies} $ t1 $ t2 =>
blanchet@37533
   341
        do_formula (flip pos) t1 #> do_formula pos t2
haftmann@39093
   342
      | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
blanchet@43606
   343
        do_term_or_formula false T t1 #> do_term_or_formula true T t2
blanchet@38810
   344
      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
blanchet@38810
   345
        $ t1 $ t2 $ t3 =>
blanchet@43600
   346
        do_formula NONE t1 #> fold (do_term_or_formula false T) [t2, t3]
blanchet@39165
   347
      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
blanchet@39165
   348
        do_quantifier (is_some pos) T t'
blanchet@39165
   349
      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
blanchet@39165
   350
        do_quantifier (pos = SOME false) T
blanchet@39165
   351
                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
blanchet@39165
   352
      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
blanchet@39165
   353
        do_quantifier (pos = SOME true) T
blanchet@39165
   354
                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
blanchet@37533
   355
      | (t0 as Const (_, @{typ bool})) $ t1 =>
blanchet@43600
   356
        do_term false t0 #> do_formula pos t1  (* theory constant *)
blanchet@43600
   357
      | _ => do_term false t
blanchet@43597
   358
  in do_formula pos end
paulson@24287
   359
paulson@24287
   360
(*Inserts a dummy "constant" referring to the theory name, so that relevance
paulson@24287
   361
  takes the given theory into account.*)
blanchet@41448
   362
fun theory_constify ({theory_const_rel_weight, theory_const_irrel_weight, ...}
blanchet@41448
   363
                     : relevance_fudge) thy_name t =
blanchet@40251
   364
  if exists (curry (op <) 0.0) [theory_const_rel_weight,
blanchet@40251
   365
                                theory_const_irrel_weight] then
blanchet@41448
   366
    Const (thy_name ^ theory_const_suffix, @{typ bool}) $ t
blanchet@37505
   367
  else
blanchet@41448
   368
    t
blanchet@41448
   369
blanchet@41448
   370
fun theory_const_prop_of fudge th =
blanchet@41448
   371
  theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
blanchet@37505
   372
paulson@24287
   373
(**** Constant / Type Frequencies ****)
paulson@24287
   374
blanchet@38982
   375
(* A two-dimensional symbol table counts frequencies of constants. It's keyed
blanchet@38982
   376
   first by constant name and second by its list of type instantiations. For the
blanchet@39058
   377
   latter, we need a linear ordering on "pattern list". *)
blanchet@37505
   378
blanchet@39058
   379
fun pattern_ord p =
blanchet@38982
   380
  case p of
blanchet@38983
   381
    (PVar, PVar) => EQUAL
blanchet@39058
   382
  | (PVar, PApp _) => LESS
blanchet@39058
   383
  | (PApp _, PVar) => GREATER
blanchet@39058
   384
  | (PApp q1, PApp q2) =>
blanchet@39058
   385
    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
blanchet@39165
   386
fun ptype_ord (PType p, PType q) =
blanchet@39165
   387
  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
paulson@24287
   388
blanchet@39165
   389
structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
paulson@24287
   390
blanchet@40445
   391
fun count_fact_consts thy fudge =
blanchet@37503
   392
  let
blanchet@39062
   393
    fun do_const const (s, T) ts =
blanchet@39062
   394
      (* Two-dimensional table update. Constant maps to types maps to count. *)
blanchet@41452
   395
      PType_Tab.map_default (rich_ptype thy const (s, T), 0) (Integer.add 1)
blanchet@39165
   396
      |> Symtab.map_default (s, PType_Tab.empty)
blanchet@39062
   397
      #> fold do_term ts
blanchet@39062
   398
    and do_term t =
blanchet@39062
   399
      case strip_comb t of
blanchet@39062
   400
        (Const x, ts) => do_const true x ts
blanchet@39062
   401
      | (Free x, ts) => do_const false x ts
blanchet@39062
   402
      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
blanchet@39062
   403
      | (_, ts) => fold do_term ts
blanchet@40251
   404
  in do_term o theory_const_prop_of fudge o snd end
paulson@24287
   405
paulson@24287
   406
paulson@24287
   407
(**** Actual Filtering Code ****)
paulson@24287
   408
blanchet@39613
   409
fun pow_int _ 0 = 1.0
blanchet@39165
   410
  | pow_int x 1 = x
blanchet@39165
   411
  | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
blanchet@39165
   412
paulson@24287
   413
(*The frequency of a constant is the sum of those of all instances of its type.*)
blanchet@39059
   414
fun pconst_freq match const_tab (c, ps) =
blanchet@39165
   415
  PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
blanchet@39165
   416
                 (the (Symtab.lookup const_tab c)) 0
blanchet@38925
   417
paulson@24287
   418
blanchet@38331
   419
(* A surprising number of theorems contain only a few significant constants.
blanchet@38331
   420
   These include all induction rules, and other general theorems. *)
blanchet@37503
   421
blanchet@37503
   422
(* "log" seems best in practice. A constant function of one ignores the constant
blanchet@39164
   423
   frequencies. Rare constants give more points if they are relevant than less
blanchet@39164
   424
   rare ones. *)
blanchet@39613
   425
fun rel_weight_for _ freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
blanchet@39164
   426
blanchet@39164
   427
(* Irrelevant constants are treated differently. We associate lower penalties to
blanchet@39164
   428
   very rare constants and very common ones -- the former because they can't
blanchet@39164
   429
   lead to the inclusion of too many new facts, and the latter because they are
blanchet@39164
   430
   so common as to be of little interest. *)
blanchet@40251
   431
fun irrel_weight_for ({worse_irrel_freq, higher_order_irrel_weight, ...}
blanchet@40251
   432
                      : relevance_fudge) order freq =
blanchet@40251
   433
  let val (k, x) = worse_irrel_freq |> `Real.ceil in
blanchet@39165
   434
    (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
blanchet@39165
   435
     else rel_weight_for order freq / rel_weight_for order k)
blanchet@40251
   436
    * pow_int higher_order_irrel_weight (order - 1)
blanchet@39164
   437
  end
blanchet@37503
   438
blanchet@42661
   439
fun multiplier_for_const_name local_const_multiplier s =
blanchet@42661
   440
  if String.isSubstring "." s then 1.0 else local_const_multiplier
blanchet@42661
   441
blanchet@37503
   442
(* Computes a constant's weight, as determined by its frequency. *)
blanchet@42661
   443
fun generic_pconst_weight local_const_multiplier abs_weight skolem_weight
blanchet@43600
   444
                          theory_const_weight chained_const_weight weight_for f
blanchet@43600
   445
                          const_tab chained_const_tab (c as (s, PType (m, _))) =
blanchet@42661
   446
  if s = abs_name then
blanchet@42661
   447
    abs_weight
blanchet@42661
   448
  else if String.isPrefix skolem_prefix s then
blanchet@42661
   449
    skolem_weight
blanchet@42661
   450
  else if String.isSuffix theory_const_suffix s then
blanchet@42661
   451
    theory_const_weight
blanchet@42661
   452
  else
blanchet@42661
   453
    multiplier_for_const_name local_const_multiplier s
blanchet@42661
   454
    * weight_for m (pconst_freq (match_ptype o f) const_tab c)
blanchet@43600
   455
    |> (if chained_const_weight < 1.0 andalso
blanchet@43600
   456
           pconst_hyper_mem I chained_const_tab c then
blanchet@43600
   457
          curry (op *) chained_const_weight
blanchet@43600
   458
        else
blanchet@43600
   459
          I)
blanchet@39056
   460
blanchet@42661
   461
fun rel_pconst_weight ({local_const_multiplier, abs_rel_weight,
blanchet@42661
   462
                        theory_const_rel_weight, ...} : relevance_fudge)
blanchet@42661
   463
                      const_tab =
blanchet@42661
   464
  generic_pconst_weight local_const_multiplier abs_rel_weight 0.0
blanchet@43600
   465
                        theory_const_rel_weight 0.0 rel_weight_for I const_tab
blanchet@43600
   466
                        Symtab.empty
blanchet@43600
   467
blanchet@42661
   468
fun irrel_pconst_weight (fudge as {local_const_multiplier, abs_irrel_weight,
blanchet@42661
   469
                                   skolem_irrel_weight,
blanchet@43600
   470
                                   theory_const_irrel_weight,
blanchet@43600
   471
                                   chained_const_irrel_weight, ...})
blanchet@43600
   472
                        const_tab chained_const_tab =
blanchet@42661
   473
  generic_pconst_weight local_const_multiplier abs_irrel_weight
blanchet@42661
   474
                        skolem_irrel_weight theory_const_irrel_weight
blanchet@43600
   475
                        chained_const_irrel_weight (irrel_weight_for fudge) swap
blanchet@43600
   476
                        const_tab chained_const_tab
paulson@24287
   477
blanchet@40251
   478
fun locality_bonus (_ : relevance_fudge) General = 0.0
blanchet@40251
   479
  | locality_bonus {intro_bonus, ...} Intro = intro_bonus
blanchet@40251
   480
  | locality_bonus {elim_bonus, ...} Elim = elim_bonus
blanchet@40251
   481
  | locality_bonus {simp_bonus, ...} Simp = simp_bonus
blanchet@40251
   482
  | locality_bonus {local_bonus, ...} Local = local_bonus
blanchet@40251
   483
  | locality_bonus {assum_bonus, ...} Assum = assum_bonus
blanchet@40251
   484
  | locality_bonus {chained_bonus, ...} Chained = chained_bonus
blanchet@38990
   485
blanchet@40645
   486
fun is_odd_const_name s =
blanchet@40645
   487
  s = abs_name orelse String.isPrefix skolem_prefix s orelse
blanchet@40645
   488
  String.isSuffix theory_const_suffix s
blanchet@40645
   489
blanchet@43597
   490
fun fact_weight fudge loc const_tab relevant_consts chained_consts fact_consts =
blanchet@40445
   491
  case fact_consts |> List.partition (pconst_hyper_mem I relevant_consts)
blanchet@40445
   492
                   ||> filter_out (pconst_hyper_mem swap relevant_consts) of
blanchet@39062
   493
    ([], _) => 0.0
blanchet@38983
   494
  | (rel, irrel) =>
blanchet@40645
   495
    if forall (forall (is_odd_const_name o fst)) [rel, irrel] then
blanchet@40617
   496
      0.0
blanchet@40617
   497
    else
blanchet@40617
   498
      let
blanchet@43600
   499
        val irrel = irrel |> filter_out (pconst_mem swap rel)
blanchet@40617
   500
        val rel_weight =
blanchet@40617
   501
          0.0 |> fold (curry (op +) o rel_pconst_weight fudge const_tab) rel
blanchet@40617
   502
        val irrel_weight =
blanchet@40617
   503
          ~ (locality_bonus fudge loc)
blanchet@43600
   504
          |> fold (curry (op +)
blanchet@43600
   505
                   o irrel_pconst_weight fudge const_tab chained_consts) irrel
blanchet@40617
   506
        val res = rel_weight / (rel_weight + irrel_weight)
blanchet@40617
   507
      in if Real.isFinite res then res else 0.0 end
blanchet@38986
   508
blanchet@40615
   509
fun pconsts_in_fact thy is_built_in_const t =
blanchet@39060
   510
  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
blanchet@43597
   511
              (Symtab.empty |> add_pconsts_in_term thy is_built_in_const true
blanchet@43597
   512
                                                   (SOME true) t) []
blanchet@43594
   513
blanchet@40615
   514
fun pair_consts_fact thy is_built_in_const fudge fact =
blanchet@40445
   515
  case fact |> snd |> theory_const_prop_of fudge
blanchet@40615
   516
            |> pconsts_in_fact thy is_built_in_const of
blanchet@39062
   517
    [] => NONE
blanchet@40445
   518
  | consts => SOME ((fact, consts), NONE)
paulson@24287
   519
blanchet@42639
   520
val const_names_in_fact = map fst ooo pconsts_in_fact
blanchet@42639
   521
blanchet@38938
   522
type annotated_thm =
blanchet@39165
   523
  (((unit -> string) * locality) * thm) * (string * ptype) list
blanchet@37505
   524
blanchet@43517
   525
fun take_most_relevant ctxt max_relevant remaining_max
blanchet@43593
   526
        ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
blanchet@40251
   527
        (candidates : (annotated_thm * real) list) =
blanchet@38983
   528
  let
blanchet@38986
   529
    val max_imperfect =
blanchet@40251
   530
      Real.ceil (Math.pow (max_imperfect,
blanchet@39130
   531
                    Math.pow (Real.fromInt remaining_max
blanchet@40251
   532
                              / Real.fromInt max_relevant, max_imperfect_exp)))
blanchet@38986
   533
    val (perfect, imperfect) =
blanchet@39115
   534
      candidates |> sort (Real.compare o swap o pairself snd)
blanchet@39115
   535
                 |> take_prefix (fn (_, w) => w > 0.99999)
blanchet@38986
   536
    val ((accepts, more_rejects), rejects) =
blanchet@38986
   537
      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
blanchet@38983
   538
  in
blanchet@43517
   539
    trace_msg ctxt (fn () =>
wenzelm@41739
   540
        "Actually passed (" ^ string_of_int (length accepts) ^ " of " ^
wenzelm@41739
   541
        string_of_int (length candidates) ^ "): " ^
blanchet@39115
   542
        (accepts |> map (fn ((((name, _), _), _), weight) =>
blanchet@38991
   543
                            name () ^ " [" ^ Real.toString weight ^ "]")
blanchet@38984
   544
                 |> commas));
blanchet@38986
   545
    (accepts, more_rejects @ rejects)
blanchet@38983
   546
  end
paulson@24287
   547
blanchet@40615
   548
fun if_empty_replace_with_locality thy is_built_in_const facts loc tab =
blanchet@39054
   549
  if Symtab.is_empty tab then
blanchet@43597
   550
    Symtab.empty
blanchet@43597
   551
    |> fold (add_pconsts_in_term thy is_built_in_const false (SOME false))
blanchet@43597
   552
            (map_filter (fn ((_, loc'), th) =>
blanchet@43597
   553
                            if loc' = loc then SOME (prop_of th) else NONE)
blanchet@43597
   554
                        facts)
blanchet@39054
   555
  else
blanchet@39054
   556
    tab
blanchet@39054
   557
blanchet@43572
   558
fun consider_arities is_built_in_const th =
blanchet@41406
   559
  let
blanchet@41406
   560
    fun aux _ _ NONE = NONE
blanchet@41406
   561
      | aux t args (SOME tab) =
blanchet@41406
   562
        case t of
blanchet@41406
   563
          t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 []
blanchet@41406
   564
        | Const (x as (s, _)) =>
blanchet@41584
   565
          (if is_built_in_const x args |> fst then
blanchet@41406
   566
             SOME tab
blanchet@41406
   567
           else case Symtab.lookup tab s of
blanchet@41406
   568
             NONE => SOME (Symtab.update (s, length args) tab)
blanchet@41406
   569
           | SOME n => if n = length args then SOME tab else NONE)
blanchet@41406
   570
        | _ => SOME tab
blanchet@41406
   571
  in aux (prop_of th) [] end
blanchet@41406
   572
blanchet@43572
   573
(* FIXME: This is currently only useful for polymorphic type systems. *)
blanchet@43572
   574
fun could_benefit_from_ext is_built_in_const facts =
blanchet@43572
   575
  fold (consider_arities is_built_in_const o snd) facts (SOME Symtab.empty)
blanchet@41406
   576
  |> is_none
blanchet@41406
   577
blanchet@40615
   578
fun relevance_filter ctxt threshold0 decay max_relevant is_built_in_const
blanchet@40251
   579
        (fudge as {threshold_divisor, ridiculous_threshold, ...})
blanchet@43597
   580
        ({add, del, ...} : relevance_override) facts chained_ts hyp_ts concl_t =
blanchet@38978
   581
  let
wenzelm@43232
   582
    val thy = Proof_Context.theory_of ctxt
blanchet@40445
   583
    val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
blanchet@43597
   584
    val add_pconsts = add_pconsts_in_term thy is_built_in_const false o SOME
blanchet@43597
   585
    val chained_const_tab = Symtab.empty |> fold (add_pconsts true) chained_ts
blanchet@39054
   586
    val goal_const_tab =
blanchet@43597
   587
      Symtab.empty |> fold (add_pconsts true) hyp_ts
blanchet@43597
   588
                   |> add_pconsts false concl_t
blanchet@43597
   589
      |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab)
blanchet@40615
   590
      |> fold (if_empty_replace_with_locality thy is_built_in_const facts)
blanchet@39237
   591
              [Chained, Assum, Local]
blanchet@39256
   592
    val add_ths = Attrib.eval_thms ctxt add
blanchet@39256
   593
    val del_ths = Attrib.eval_thms ctxt del
blanchet@43798
   594
    val facts = facts |> filter_out (member Thm.eq_thm_prop del_ths o snd)
blanchet@38986
   595
    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
blanchet@38978
   596
      let
blanchet@40432
   597
        fun relevant [] _ [] =
blanchet@38986
   598
            (* Nothing has been added this iteration. *)
blanchet@40251
   599
            if j = 0 andalso threshold >= ridiculous_threshold then
blanchet@38986
   600
              (* First iteration? Try again. *)
blanchet@40251
   601
              iter 0 max_relevant (threshold / threshold_divisor) rel_const_tab
blanchet@38986
   602
                   hopeless hopeful
blanchet@38983
   603
            else
blanchet@40432
   604
              []
blanchet@39115
   605
          | relevant candidates rejects [] =
blanchet@38978
   606
            let
blanchet@38986
   607
              val (accepts, more_rejects) =
blanchet@43517
   608
                take_most_relevant ctxt max_relevant remaining_max fudge
blanchet@43517
   609
                                   candidates
blanchet@38978
   610
              val rel_const_tab' =
blanchet@38984
   611
                rel_const_tab
blanchet@41314
   612
                |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
blanchet@38983
   613
              fun is_dirty (c, _) =
blanchet@38983
   614
                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
blanchet@38984
   615
              val (hopeful_rejects, hopeless_rejects) =
blanchet@38984
   616
                 (rejects @ hopeless, ([], []))
blanchet@38984
   617
                 |-> fold (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   618
                              if exists is_dirty consts then
blanchet@38984
   619
                                apfst (cons (ax, NONE))
blanchet@38984
   620
                              else
blanchet@38984
   621
                                apsnd (cons (ax, old_weight)))
blanchet@38984
   622
                 |>> append (more_rejects
blanchet@38984
   623
                             |> map (fn (ax as (_, consts), old_weight) =>
blanchet@38984
   624
                                        (ax, if exists is_dirty consts then NONE
blanchet@38984
   625
                                             else SOME old_weight)))
blanchet@38986
   626
              val threshold =
blanchet@39057
   627
                1.0 - (1.0 - threshold)
blanchet@39057
   628
                      * Math.pow (decay, Real.fromInt (length accepts))
blanchet@38986
   629
              val remaining_max = remaining_max - length accepts
blanchet@38978
   630
            in
blanchet@43517
   631
              trace_msg ctxt (fn () => "New or updated constants: " ^
blanchet@38983
   632
                  commas (rel_const_tab' |> Symtab.dest
blanchet@39057
   633
                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
blanchet@39062
   634
                          |> map string_for_hyper_pconst));
blanchet@38984
   635
              map (fst o fst) accepts @
blanchet@38986
   636
              (if remaining_max = 0 then
blanchet@40432
   637
                 []
blanchet@38984
   638
               else
blanchet@38986
   639
                 iter (j + 1) remaining_max threshold rel_const_tab'
blanchet@38986
   640
                      hopeless_rejects hopeful_rejects)
blanchet@38978
   641
            end
blanchet@39115
   642
          | relevant candidates rejects
blanchet@40445
   643
                     (((ax as (((_, loc), _), fact_consts)), cached_weight)
blanchet@38986
   644
                      :: hopeful) =
blanchet@38978
   645
            let
blanchet@38978
   646
              val weight =
blanchet@38978
   647
                case cached_weight of
blanchet@38978
   648
                  SOME w => w
blanchet@40445
   649
                | NONE => fact_weight fudge loc const_tab rel_const_tab
blanchet@43597
   650
                                      chained_const_tab fact_consts
blanchet@38978
   651
            in
blanchet@38980
   652
              if weight >= threshold then
blanchet@39115
   653
                relevant ((ax, weight) :: candidates) rejects hopeful
blanchet@38978
   654
              else
blanchet@39115
   655
                relevant candidates ((ax, weight) :: rejects) hopeful
blanchet@38978
   656
            end
blanchet@38978
   657
        in
blanchet@43517
   658
          trace_msg ctxt (fn () =>
blanchet@38983
   659
              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
blanchet@38983
   660
              Real.toString threshold ^ ", constants: " ^
blanchet@38983
   661
              commas (rel_const_tab |> Symtab.dest
blanchet@38983
   662
                      |> filter (curry (op <>) [] o snd)
blanchet@39062
   663
                      |> map string_for_hyper_pconst));
blanchet@39115
   664
          relevant [] [] hopeful
blanchet@38978
   665
        end
blanchet@43697
   666
    fun prepend_facts ths accepts =
blanchet@43798
   667
      ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@43798
   668
       (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@40635
   669
      |> take max_relevant
blanchet@43907
   670
    fun append_to_facts accepts ths =
blanchet@43798
   671
      let val add = facts |> filter (member Thm.eq_thm_prop ths o snd) in
blanchet@43798
   672
        (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)
blanchet@43697
   673
                 |> take (max_relevant - length add)) @
blanchet@43697
   674
        add
blanchet@43697
   675
      end
blanchet@43606
   676
    fun uses_const s t =
blanchet@43606
   677
      fold_aterms (curry (fn (Const (s', _), false) => s' = s | (_, b) => b)) t
blanchet@43606
   678
                  false
blanchet@43907
   679
    fun add_set_const_thms accepts (s, ths) =
blanchet@43907
   680
      if exists (uses_const s o prop_of o snd) accepts orelse
blanchet@43907
   681
         exists (uses_const s) (concl_t :: hyp_ts) then
blanchet@43907
   682
        append ths
blanchet@43907
   683
      else
blanchet@43907
   684
        I
blanchet@43907
   685
    fun append_special_facts accepts =
blanchet@43907
   686
       [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
blanchet@43907
   687
          |> fold (add_set_const_thms accepts) set_consts
blanchet@43907
   688
          |> append_to_facts accepts
blanchet@43907
   689
blanchet@38978
   690
  in
blanchet@40615
   691
    facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
blanchet@40445
   692
          |> iter 0 max_relevant threshold0 goal_const_tab []
blanchet@43697
   693
          |> not (null add_ths) ? prepend_facts add_ths
blanchet@43907
   694
          |> append_special_facts
blanchet@43517
   695
          |> tap (fn accepts => trace_msg ctxt (fn () =>
wenzelm@41739
   696
                      "Total relevant: " ^ string_of_int (length accepts)))
blanchet@38978
   697
  end
paulson@24287
   698
blanchet@38983
   699
paulson@24287
   700
(***************************************************************)
mengj@19768
   701
(* Retrieving and filtering lemmas                             *)
mengj@19768
   702
(***************************************************************)
mengj@19768
   703
paulson@33022
   704
(*** retrieve lemmas and filter them ***)
mengj@19768
   705
paulson@20757
   706
(*Reject theorems with names like "List.filter.filter_list_def" or
paulson@21690
   707
  "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
paulson@20757
   708
fun is_package_def a =
blanchet@40446
   709
  let val names = Long_Name.explode a in
blanchet@40446
   710
    (length names > 2 andalso not (hd names = "local") andalso
blanchet@40446
   711
     String.isSuffix "_def" a) orelse String.isSuffix "_defs" a
blanchet@40446
   712
  end
paulson@20757
   713
blanchet@44086
   714
fun uniquify xs =
blanchet@44086
   715
  Termtab.fold (cons o snd)
blanchet@44086
   716
               (fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty) []
mengj@19768
   717
blanchet@37626
   718
(* FIXME: put other record thms here, or declare as "no_atp" *)
blanchet@43593
   719
fun multi_base_blacklist ctxt =
blanchet@41447
   720
  ["defs", "select_defs", "update_defs", "split", "splits", "split_asm",
blanchet@41447
   721
   "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy", "case_cong",
blanchet@41447
   722
   "weak_case_cong"]
blanchet@43593
   723
  |> not (Config.get ctxt instantiate_inducts) ? append ["induct", "inducts"]
blanchet@38921
   724
  |> map (prefix ".")
blanchet@37626
   725
blanchet@37626
   726
val max_lambda_nesting = 3
blanchet@37626
   727
blanchet@37626
   728
fun term_has_too_many_lambdas max (t1 $ t2) =
blanchet@37626
   729
    exists (term_has_too_many_lambdas max) [t1, t2]
blanchet@37626
   730
  | term_has_too_many_lambdas max (Abs (_, _, t)) =
blanchet@37626
   731
    max = 0 orelse term_has_too_many_lambdas (max - 1) t
blanchet@37626
   732
  | term_has_too_many_lambdas _ _ = false
blanchet@37626
   733
blanchet@37626
   734
(* Don't count nested lambdas at the level of formulas, since they are
blanchet@37626
   735
   quantifiers. *)
blanchet@37626
   736
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
blanchet@37626
   737
    formula_has_too_many_lambdas (T :: Ts) t
blanchet@37626
   738
  | formula_has_too_many_lambdas Ts t =
blanchet@41521
   739
    if member (op =) [HOLogic.boolT, propT] (fastype_of1 (Ts, t)) then
blanchet@37626
   740
      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
blanchet@37626
   741
    else
blanchet@37626
   742
      term_has_too_many_lambdas max_lambda_nesting t
blanchet@37626
   743
blanchet@38931
   744
(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
blanchet@37626
   745
   was 11. *)
blanchet@37626
   746
val max_apply_depth = 15
blanchet@37626
   747
blanchet@37626
   748
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
blanchet@37626
   749
  | apply_depth (Abs (_, _, t)) = apply_depth t
blanchet@37626
   750
  | apply_depth _ = 0
blanchet@37626
   751
blanchet@37626
   752
fun is_formula_too_complex t =
blanchet@38331
   753
  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
blanchet@37626
   754
blanchet@40127
   755
(* FIXME: Extend to "Meson" and "Metis" *)
blanchet@37539
   756
val exists_sledgehammer_const =
blanchet@37626
   757
  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
blanchet@37626
   758
blanchet@39130
   759
(* FIXME: make more reliable *)
blanchet@39130
   760
val exists_low_level_class_const =
blanchet@39130
   761
  exists_Const (fn (s, _) =>
blanchet@39130
   762
     String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
blanchet@39130
   763
blanchet@39056
   764
fun is_metastrange_theorem th =
blanchet@37626
   765
  case head_of (concl_of th) of
blanchet@37626
   766
      Const (a, _) => (a <> @{const_name Trueprop} andalso
blanchet@37626
   767
                       a <> @{const_name "=="})
blanchet@37626
   768
    | _ => false
blanchet@37626
   769
blanchet@39056
   770
fun is_that_fact th =
blanchet@39056
   771
  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
blanchet@39056
   772
  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
blanchet@39056
   773
                           | _ => false) (prop_of th)
blanchet@39056
   774
blanchet@37626
   775
val type_has_top_sort =
blanchet@37626
   776
  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
blanchet@37626
   777
blanchet@38331
   778
(**** Predicates to detect unwanted facts (prolific or likely to cause
blanchet@37347
   779
      unsoundness) ****)
paulson@21470
   780
blanchet@43793
   781
fun is_theorem_bad_for_atps is_appropriate_prop thm =
blanchet@38850
   782
  let val t = prop_of thm in
blanchet@43793
   783
    not (is_appropriate_prop t) orelse is_formula_too_complex t orelse
blanchet@43785
   784
    exists_type type_has_top_sort t orelse exists_sledgehammer_const t orelse
blanchet@43785
   785
    exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
blanchet@43785
   786
    is_that_fact thm
blanchet@38850
   787
  end
blanchet@38850
   788
blanchet@43512
   789
fun meta_equify (@{const Trueprop}
blanchet@43512
   790
                 $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2)) =
blanchet@43512
   791
    Const (@{const_name "=="}, T --> T --> @{typ prop}) $ t1 $ t2
blanchet@43512
   792
  | meta_equify t = t
blanchet@43512
   793
blanchet@43512
   794
val normalize_simp_prop =
blanchet@43512
   795
  meta_equify
blanchet@43512
   796
  #> map_aterms (fn Var ((s, _), T) => Var ((s, 0), T) | t => t)
blanchet@43512
   797
  #> map_types (map_type_tvar (fn ((s, _), S) => TVar ((s, 0), S)))
blanchet@43512
   798
blanchet@44086
   799
fun clasimpset_rule_table_of ctxt =
blanchet@39163
   800
  let
blanchet@44086
   801
    fun add loc g f = fold (Termtab.update o rpair loc o g o prop_of o f)
wenzelm@43683
   802
    val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> Classical.rep_cs
wenzelm@43681
   803
    val intros = Item_Net.content safeIs @ Item_Net.content hazIs
wenzelm@43681
   804
    val elims = map Classical.classical_rule (Item_Net.content safeEs @ Item_Net.content hazEs)
blanchet@43512
   805
    val simps = ctxt |> simpset_of |> dest_ss |> #simps
blanchet@43512
   806
  in
blanchet@44086
   807
    Termtab.empty |> add Intro I I intros
blanchet@44086
   808
                  |> add Elim I I elims
blanchet@44086
   809
                  |> add Simp I snd simps
blanchet@44086
   810
                  |> add Simp normalize_simp_prop snd simps
blanchet@43512
   811
  end
blanchet@39163
   812
blanchet@43793
   813
fun all_facts ctxt reserved really_all is_appropriate_prop add_ths chained_ths =
blanchet@38850
   814
  let
wenzelm@43232
   815
    val thy = Proof_Context.theory_of ctxt
wenzelm@39814
   816
    val global_facts = Global_Theory.facts_of thy
wenzelm@43232
   817
    val local_facts = Proof_Context.facts_of ctxt
blanchet@38882
   818
    val named_locals = local_facts |> Facts.dest_static []
blanchet@39237
   819
    val assms = Assumption.all_assms_of ctxt
blanchet@39237
   820
    fun is_assum th = exists (fn ct => prop_of th aconv term_of ct) assms
blanchet@43798
   821
    val is_chained = member Thm.eq_thm_prop chained_ths
blanchet@44086
   822
    val clasimpset_table = clasimpset_rule_table_of ctxt
blanchet@43512
   823
    fun locality_of_theorem global th =
blanchet@43512
   824
      if is_chained th then
blanchet@43512
   825
        Chained
blanchet@43512
   826
      else if global then
blanchet@44086
   827
        Termtab.lookup clasimpset_table (prop_of th) |> the_default General
blanchet@39163
   828
      else
blanchet@43512
   829
        if is_assum th then Assum else Local
blanchet@38977
   830
    fun is_good_unnamed_local th =
blanchet@39055
   831
      not (Thm.has_name_hint th) andalso
blanchet@43798
   832
      forall (fn (_, ths) => not (member Thm.eq_thm_prop ths th)) named_locals
blanchet@38882
   833
    val unnamed_locals =
blanchet@43798
   834
      union Thm.eq_thm_prop (Facts.props local_facts) chained_ths
blanchet@39055
   835
      |> filter is_good_unnamed_local |> map (pair "" o single)
blanchet@38850
   836
    val full_space =
blanchet@38977
   837
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
blanchet@38991
   838
    fun add_facts global foldx facts =
blanchet@38938
   839
      foldx (fn (name0, ths) =>
blanchet@42860
   840
        if not really_all andalso name0 <> "" andalso
blanchet@43798
   841
           forall (not o member Thm.eq_thm_prop add_ths) ths andalso
blanchet@38938
   842
           (Facts.is_concealed facts name0 orelse
blanchet@43593
   843
            (not (Config.get ctxt ignore_no_atp) andalso
blanchet@43593
   844
             is_package_def name0) orelse
blanchet@43593
   845
            exists (fn s => String.isSuffix s name0)
blanchet@43593
   846
                   (multi_base_blacklist ctxt) orelse
blanchet@38938
   847
            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
blanchet@38850
   848
          I
blanchet@38850
   849
        else
blanchet@38850
   850
          let
blanchet@38938
   851
            val multi = length ths > 1
blanchet@41527
   852
            val backquote_thm =
blanchet@41527
   853
              backquote o string_for_term ctxt o close_form o prop_of
blanchet@38938
   854
            fun check_thms a =
wenzelm@43232
   855
              case try (Proof_Context.get_thms ctxt) a of
blanchet@38938
   856
                NONE => false
blanchet@43798
   857
              | SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')
blanchet@38850
   858
          in
blanchet@38938
   859
            pair 1
blanchet@44086
   860
            #> fold (fn th => fn (j, (multis, unis)) =>
blanchet@43512
   861
                        (j + 1,
blanchet@43798
   862
                         if not (member Thm.eq_thm_prop add_ths th) andalso
blanchet@43793
   863
                            is_theorem_bad_for_atps is_appropriate_prop th then
blanchet@44086
   864
                           (multis, unis)
blanchet@43512
   865
                         else
blanchet@44086
   866
                           let
blanchet@44086
   867
                             val new =
blanchet@44086
   868
                               (((fn () =>
blanchet@43512
   869
                                 if name0 = "" then
blanchet@43512
   870
                                   th |> backquote_thm
blanchet@43512
   871
                                 else
blanchet@43512
   872
                                   [Facts.extern ctxt facts name0,
blanchet@43512
   873
                                    Name_Space.extern ctxt full_space name0,
blanchet@43512
   874
                                    name0]
blanchet@43512
   875
                                   |> find_first check_thms
blanchet@43512
   876
                                   |> (fn SOME name =>
blanchet@43512
   877
                                          make_name reserved multi j name
blanchet@43512
   878
                                        | NONE => "")),
blanchet@44086
   879
                                locality_of_theorem global th), th)
blanchet@44086
   880
                           in
blanchet@44086
   881
                             if multi then (new :: multis, unis)
blanchet@44086
   882
                             else (multis, new :: unis)
blanchet@44086
   883
                           end)) ths
blanchet@38938
   884
            #> snd
blanchet@38850
   885
          end)
blanchet@38882
   886
  in
blanchet@44086
   887
    (* The single-name theorems go after the multiple-name ones, so that single
blanchet@44086
   888
       names are preferred when both are available. *)
blanchet@44086
   889
    ([], []) |> add_facts false fold local_facts (unnamed_locals @ named_locals)
blanchet@44086
   890
             |> add_facts true Facts.fold_static global_facts global_facts
blanchet@44086
   891
             |> op @
blanchet@38882
   892
  end
blanchet@38850
   893
blanchet@41447
   894
fun external_frees t =
blanchet@41447
   895
  [] |> Term.add_frees t |> filter_out (can Name.dest_internal o fst)
blanchet@41447
   896
blanchet@43793
   897
fun relevant_facts ctxt (threshold0, threshold1) max_relevant
blanchet@43793
   898
        is_appropriate_prop is_built_in_const fudge
blanchet@43793
   899
        (override as {add, only, ...}) chained_ths hyp_ts concl_t =
blanchet@37534
   900
  let
wenzelm@43232
   901
    val thy = Proof_Context.theory_of ctxt
blanchet@39057
   902
    val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
blanchet@39057
   903
                          1.0 / Real.fromInt (max_relevant + 1))
blanchet@39256
   904
    val add_ths = Attrib.eval_thms ctxt add
blanchet@38935
   905
    val reserved = reserved_isar_keyword_table ()
blanchet@41447
   906
    val ind_stmt =
blanchet@41447
   907
      Logic.list_implies (hyp_ts |> filter_out (null o external_frees), concl_t)
blanchet@41459
   908
      |> Object_Logic.atomize_term thy
blanchet@41447
   909
    val ind_stmt_xs = external_frees ind_stmt
blanchet@40445
   910
    val facts =
blanchet@38938
   911
      (if only then
blanchet@44086
   912
         maps (map (fn ((name, loc), th) => ((K name, loc), th))
blanchet@40446
   913
               o fact_from_ref ctxt reserved chained_ths) add
blanchet@38938
   914
       else
blanchet@43793
   915
         all_facts ctxt reserved false is_appropriate_prop add_ths chained_ths)
blanchet@43593
   916
      |> Config.get ctxt instantiate_inducts
blanchet@41521
   917
         ? maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs)
blanchet@44086
   918
      |> (not (Config.get ctxt ignore_no_atp) andalso not only)
blanchet@44086
   919
         ? filter_out (No_ATPs.member ctxt o snd)
blanchet@39163
   920
      |> uniquify
blanchet@37534
   921
  in
blanchet@43517
   922
    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^
blanchet@43517
   923
                             " facts");
blanchet@39612
   924
    (if only orelse threshold1 < 0.0 then
blanchet@40445
   925
       facts
blanchet@39612
   926
     else if threshold0 > 1.0 orelse threshold0 > threshold1 orelse
blanchet@39612
   927
             max_relevant = 0 then
blanchet@38978
   928
       []
blanchet@38978
   929
     else
blanchet@43597
   930
       relevance_filter ctxt threshold0 decay max_relevant is_built_in_const
blanchet@43597
   931
           fudge override facts (chained_ths |> map prop_of) hyp_ts
blanchet@43597
   932
           (concl_t |> theory_constify fudge (Context.theory_name thy)))
blanchet@39057
   933
    |> map (apfst (apfst (fn f => f ())))
blanchet@37534
   934
  end
immler@30536
   935
paulson@15347
   936
end;