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