src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML
changeset 48250 1065c307fafe
parent 48248 b6eb45a52c28
equal deleted inserted replaced
48249:2bd242c56c90 48250:1065c307fafe
     5 Sledgehammer's hybrid relevance filter.
     5 Sledgehammer's hybrid relevance filter.
     6 *)
     6 *)
     7 
     7 
     8 signature SLEDGEHAMMER_FILTER =
     8 signature SLEDGEHAMMER_FILTER =
     9 sig
     9 sig
    10   type status = ATP_Problem_Generate.status
       
    11   type stature = ATP_Problem_Generate.stature
    10   type stature = ATP_Problem_Generate.stature
    12   type relevance_fudge = Sledgehammer_Filter_Iter.relevance_fudge
    11   type relevance_fudge = Sledgehammer_Filter_Iter.relevance_fudge
    13   type relevance_override = Sledgehammer_Filter_Iter.relevance_override
    12   type relevance_override = Sledgehammer_Filter_Iter.relevance_override
    14 
    13 
    15   val no_relevance_override : relevance_override
       
    16   val fact_from_ref :
       
    17     Proof.context -> unit Symtab.table -> thm list -> status Termtab.table
       
    18     -> Facts.ref * Attrib.src list -> ((string * stature) * thm) list
       
    19   val all_facts :
       
    20     Proof.context -> bool -> 'a Symtab.table -> bool -> thm list
       
    21     -> thm list -> status Termtab.table
       
    22     -> (((unit -> string) * stature) * thm) list
       
    23   val clasimpset_rule_table_of : Proof.context -> status Termtab.table
       
    24   val maybe_instantiate_inducts :
       
    25     Proof.context -> term list -> term -> (((unit -> string) * 'a) * thm) list
       
    26     -> (((unit -> string) * 'a) * thm) list
       
    27   val maybe_filter_no_atps : Proof.context -> ('a * thm) list -> ('a * thm) list
       
    28   val nearly_all_facts :
       
    29     Proof.context -> bool -> relevance_override -> thm list -> term list -> term
       
    30     -> (((unit -> string) * stature) * thm) list
       
    31   val relevant_facts :
    14   val relevant_facts :
    32     Proof.context -> real * real -> int
    15     Proof.context -> real * real -> int
    33     -> (string * typ -> term list -> bool * term list) -> relevance_fudge
    16     -> (string * typ -> term list -> bool * term list) -> relevance_fudge
    34     -> relevance_override -> thm list -> term list -> term
    17     -> relevance_override -> thm list -> term list -> term
    35     -> (((unit -> string) * stature) * thm) list
    18     -> (((unit -> string) * stature) * thm) list
    37 end;
    20 end;
    38 
    21 
    39 structure Sledgehammer_Filter : SLEDGEHAMMER_FILTER =
    22 structure Sledgehammer_Filter : SLEDGEHAMMER_FILTER =
    40 struct
    23 struct
    41 
    24 
    42 open ATP_Problem_Generate
       
    43 open Metis_Tactic
       
    44 open Sledgehammer_Util
       
    45 open Sledgehammer_Filter_Iter
    25 open Sledgehammer_Filter_Iter
    46 
       
    47 val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
       
    48 
       
    49 val no_relevance_override = {add = [], del = [], only = false}
       
    50 
       
    51 fun needs_quoting reserved s =
       
    52   Symtab.defined reserved s orelse
       
    53   exists (not o Lexicon.is_identifier) (Long_Name.explode s)
       
    54 
       
    55 fun make_name reserved multi j name =
       
    56   (name |> needs_quoting reserved name ? quote) ^
       
    57   (if multi then "(" ^ string_of_int j ^ ")" else "")
       
    58 
       
    59 fun explode_interval _ (Facts.FromTo (i, j)) = i upto j
       
    60   | explode_interval max (Facts.From i) = i upto i + max - 1
       
    61   | explode_interval _ (Facts.Single i) = [i]
       
    62 
       
    63 val backquote =
       
    64   raw_explode #> map (fn "`" => "\\`" | s => s) #> implode #> enclose "`" "`"
       
    65 
       
    66 (* unfolding these can yield really huge terms *)
       
    67 val risky_defs = @{thms Bit0_def Bit1_def}
       
    68 
       
    69 fun is_rec_eq lhs = Term.exists_subterm (curry (op =) (head_of lhs))
       
    70 fun is_rec_def (@{const Trueprop} $ t) = is_rec_def t
       
    71   | is_rec_def (@{const ==>} $ _ $ t2) = is_rec_def t2
       
    72   | is_rec_def (Const (@{const_name "=="}, _) $ t1 $ t2) = is_rec_eq t1 t2
       
    73   | is_rec_def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) = is_rec_eq t1 t2
       
    74   | is_rec_def _ = false
       
    75 
       
    76 fun is_assum assms th = exists (fn ct => prop_of th aconv term_of ct) assms
       
    77 fun is_chained chained_ths = member Thm.eq_thm_prop chained_ths
       
    78 
       
    79 fun scope_of_thm global assms chained_ths th =
       
    80   if is_chained chained_ths th then Chained
       
    81   else if global then Global
       
    82   else if is_assum assms th then Assum
       
    83   else Local
       
    84 
       
    85 val may_be_induction =
       
    86   exists_subterm (fn Var (_, Type (@{type_name fun}, [_, T])) =>
       
    87                      body_type T = @{typ bool}
       
    88                    | _ => false)
       
    89 
       
    90 fun status_of_thm css_table name th =
       
    91   (* FIXME: use structured name *)
       
    92   if (String.isSubstring ".induct" name orelse
       
    93       String.isSubstring ".inducts" name) andalso
       
    94      may_be_induction (prop_of th) then
       
    95     Induction
       
    96   else case Termtab.lookup css_table (prop_of th) of
       
    97     SOME status => status
       
    98   | NONE => General
       
    99 
       
   100 fun stature_of_thm global assms chained_ths css_table name th =
       
   101   (scope_of_thm global assms chained_ths th, status_of_thm css_table name th)
       
   102 
       
   103 fun fact_from_ref ctxt reserved chained_ths css_table (xthm as (xref, args)) =
       
   104   let
       
   105     val ths = Attrib.eval_thms ctxt [xthm]
       
   106     val bracket =
       
   107       map (enclose "[" "]" o Pretty.str_of o Args.pretty_src ctxt) args
       
   108       |> implode
       
   109     fun nth_name j =
       
   110       case xref of
       
   111         Facts.Fact s => backquote s ^ bracket
       
   112       | Facts.Named (("", _), _) => "[" ^ bracket ^ "]"
       
   113       | Facts.Named ((name, _), NONE) =>
       
   114         make_name reserved (length ths > 1) (j + 1) name ^ bracket
       
   115       | Facts.Named ((name, _), SOME intervals) =>
       
   116         make_name reserved true
       
   117                  (nth (maps (explode_interval (length ths)) intervals) j) name ^
       
   118         bracket
       
   119   in
       
   120     (ths, (0, []))
       
   121     |-> fold (fn th => fn (j, rest) =>
       
   122                  let val name = nth_name j in
       
   123                    (j + 1, ((name, stature_of_thm false [] chained_ths
       
   124                                              css_table name th), th) :: rest)
       
   125                  end)
       
   126     |> snd
       
   127   end
       
   128 
       
   129 (*Reject theorems with names like "List.filter.filter_list_def" or
       
   130   "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
       
   131 fun is_package_def a =
       
   132   let val names = Long_Name.explode a in
       
   133     (length names > 2 andalso not (hd names = "local") andalso
       
   134      String.isSuffix "_def" a) orelse String.isSuffix "_defs" a
       
   135   end
       
   136 
       
   137 (* FIXME: put other record thms here, or declare as "no_atp" *)
       
   138 fun multi_base_blacklist ctxt ho_atp =
       
   139   ["defs", "select_defs", "update_defs", "split", "splits", "split_asm",
       
   140    "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy", "case_cong",
       
   141    "weak_case_cong"]
       
   142   |> not (ho_atp orelse (Config.get ctxt instantiate_inducts)) ?
       
   143         append ["induct", "inducts"]
       
   144   |> map (prefix ".")
       
   145 
       
   146 val max_lambda_nesting = 3 (*only applies if not ho_atp*)
       
   147 
       
   148 fun term_has_too_many_lambdas max (t1 $ t2) =
       
   149     exists (term_has_too_many_lambdas max) [t1, t2]
       
   150   | term_has_too_many_lambdas max (Abs (_, _, t)) =
       
   151     max = 0 orelse term_has_too_many_lambdas (max - 1) t
       
   152   | term_has_too_many_lambdas _ _ = false
       
   153 
       
   154 (* Don't count nested lambdas at the level of formulas, since they are
       
   155    quantifiers. *)
       
   156 fun formula_has_too_many_lambdas true _ _ = false (*i.e. ho_atp*)
       
   157   | formula_has_too_many_lambdas _ Ts (Abs (_, T, t)) =
       
   158       formula_has_too_many_lambdas false (T :: Ts) t
       
   159   | formula_has_too_many_lambdas _ Ts t =
       
   160     if member (op =) [HOLogic.boolT, propT] (fastype_of1 (Ts, t)) then
       
   161       exists (formula_has_too_many_lambdas false Ts) (#2 (strip_comb t))
       
   162     else
       
   163       term_has_too_many_lambdas max_lambda_nesting t
       
   164 
       
   165 (* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
       
   166    was 11. *)
       
   167 val max_apply_depth = 15
       
   168 
       
   169 fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
       
   170   | apply_depth (Abs (_, _, t)) = apply_depth t
       
   171   | apply_depth _ = 0
       
   172 
       
   173 fun is_formula_too_complex ho_atp t =
       
   174   apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas ho_atp [] t
       
   175 
       
   176 (* FIXME: Extend to "Meson" and "Metis" *)
       
   177 val exists_sledgehammer_const =
       
   178   exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
       
   179 
       
   180 (* FIXME: make more reliable *)
       
   181 val exists_low_level_class_const =
       
   182   exists_Const (fn (s, _) =>
       
   183      s = @{const_name equal_class.equal} orelse
       
   184      String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
       
   185 
       
   186 fun is_metastrange_theorem th =
       
   187   case head_of (concl_of th) of
       
   188     Const (s, _) => (s <> @{const_name Trueprop} andalso
       
   189                      s <> @{const_name "=="})
       
   190   | _ => false
       
   191 
       
   192 fun is_that_fact th =
       
   193   String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
       
   194   andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
       
   195                            | _ => false) (prop_of th)
       
   196 
       
   197 fun is_theorem_bad_for_atps ho_atp exporter thm =
       
   198   is_metastrange_theorem thm orelse
       
   199   (not exporter andalso
       
   200    let val t = prop_of thm in
       
   201      is_formula_too_complex ho_atp t orelse exists_type type_has_top_sort t orelse
       
   202      exists_sledgehammer_const t orelse exists_low_level_class_const t orelse
       
   203      is_that_fact thm
       
   204    end)
       
   205 
       
   206 fun string_for_term ctxt t =
       
   207   Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
       
   208                    (print_mode_value ())) (Syntax.string_of_term ctxt) t
       
   209   |> String.translate (fn c => if Char.isPrint c then str c else "")
       
   210   |> simplify_spaces
       
   211 
       
   212 (* This is a terrible hack. Free variables are sometimes coded as "M__" when
       
   213    they are displayed as "M" and we want to avoid clashes with these. But
       
   214    sometimes it's even worse: "Ma__" encodes "M". So we simply reserve all
       
   215    prefixes of all free variables. In the worse case scenario, where the fact
       
   216    won't be resolved correctly, the user can fix it manually, e.g., by naming
       
   217    the fact in question. Ideally we would need nothing of it, but backticks
       
   218    simply don't work with schematic variables. *)
       
   219 fun all_prefixes_of s =
       
   220   map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1)
       
   221 
       
   222 fun close_form t =
       
   223   (t, [] |> Term.add_free_names t |> maps all_prefixes_of)
       
   224   |> fold (fn ((s, i), T) => fn (t', taken) =>
       
   225               let val s' = singleton (Name.variant_list taken) s in
       
   226                 ((if fastype_of t' = HOLogic.boolT then HOLogic.all_const
       
   227                   else Logic.all_const) T
       
   228                  $ Abs (s', T, abstract_over (Var ((s, i), T), t')),
       
   229                  s' :: taken)
       
   230               end)
       
   231           (Term.add_vars t [] |> sort_wrt (fst o fst))
       
   232   |> fst
       
   233 
       
   234 fun all_facts ctxt ho_atp reserved exporter add_ths chained_ths css_table =
       
   235   let
       
   236     val thy = Proof_Context.theory_of ctxt
       
   237     val global_facts = Global_Theory.facts_of thy
       
   238     val local_facts = Proof_Context.facts_of ctxt
       
   239     val named_locals = local_facts |> Facts.dest_static []
       
   240     val assms = Assumption.all_assms_of ctxt
       
   241     fun is_good_unnamed_local th =
       
   242       not (Thm.has_name_hint th) andalso
       
   243       forall (fn (_, ths) => not (member Thm.eq_thm_prop ths th)) named_locals
       
   244     val unnamed_locals =
       
   245       union Thm.eq_thm_prop (Facts.props local_facts) chained_ths
       
   246       |> filter is_good_unnamed_local |> map (pair "" o single)
       
   247     val full_space =
       
   248       Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
       
   249     fun add_facts global foldx facts =
       
   250       foldx (fn (name0, ths) =>
       
   251         if not exporter andalso name0 <> "" andalso
       
   252            forall (not o member Thm.eq_thm_prop add_ths) ths andalso
       
   253            (Facts.is_concealed facts name0 orelse
       
   254             (not (Config.get ctxt ignore_no_atp) andalso
       
   255              is_package_def name0) orelse
       
   256             exists (fn s => String.isSuffix s name0)
       
   257                    (multi_base_blacklist ctxt ho_atp)) then
       
   258           I
       
   259         else
       
   260           let
       
   261             val multi = length ths > 1
       
   262             val backquote_thm =
       
   263               backquote o string_for_term ctxt o close_form o prop_of
       
   264             fun check_thms a =
       
   265               case try (Proof_Context.get_thms ctxt) a of
       
   266                 NONE => false
       
   267               | SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')
       
   268           in
       
   269             pair 1
       
   270             #> fold (fn th => fn (j, (multis, unis)) =>
       
   271                         (j + 1,
       
   272                          if not (member Thm.eq_thm_prop add_ths th) andalso
       
   273                             is_theorem_bad_for_atps ho_atp exporter th then
       
   274                            (multis, unis)
       
   275                          else
       
   276                            let
       
   277                              val new =
       
   278                                (((fn () =>
       
   279                                  if name0 = "" then
       
   280                                    th |> backquote_thm
       
   281                                  else
       
   282                                    [Facts.extern ctxt facts name0,
       
   283                                     Name_Space.extern ctxt full_space name0,
       
   284                                     name0]
       
   285                                    |> find_first check_thms
       
   286                                    |> (fn SOME name =>
       
   287                                           make_name reserved multi j name
       
   288                                         | NONE => "")),
       
   289                                 stature_of_thm global assms chained_ths
       
   290                                                css_table name0 th), th)
       
   291                            in
       
   292                              if multi then (new :: multis, unis)
       
   293                              else (multis, new :: unis)
       
   294                            end)) ths
       
   295             #> snd
       
   296           end)
       
   297   in
       
   298     (* The single-name theorems go after the multiple-name ones, so that single
       
   299        names are preferred when both are available. *)
       
   300     ([], []) |> add_facts false fold local_facts (unnamed_locals @ named_locals)
       
   301              |> add_facts true Facts.fold_static global_facts global_facts
       
   302              |> op @
       
   303   end
       
   304 
       
   305 fun clasimpset_rule_table_of ctxt =
       
   306   let
       
   307     val thy = Proof_Context.theory_of ctxt
       
   308     val atomize = HOLogic.mk_Trueprop o Object_Logic.atomize_term thy
       
   309     fun add stature normalizers get_th =
       
   310       fold (fn rule =>
       
   311                let
       
   312                  val th = rule |> get_th
       
   313                  val t =
       
   314                    th |> Thm.maxidx_of th > 0 ? zero_var_indexes |> prop_of
       
   315                in
       
   316                  fold (fn normalize => Termtab.update (normalize t, stature))
       
   317                       (I :: normalizers)
       
   318                end)
       
   319     val {safeIs, (* safeEs, *) hazIs, (* hazEs, *) ...} =
       
   320       ctxt |> claset_of |> Classical.rep_cs
       
   321     val intros = Item_Net.content safeIs @ Item_Net.content hazIs
       
   322 (* Add once it is used:
       
   323     val elims =
       
   324       Item_Net.content safeEs @ Item_Net.content hazEs
       
   325       |> map Classical.classical_rule
       
   326 *)
       
   327     val simps = ctxt |> simpset_of |> dest_ss |> #simps
       
   328     val specs = ctxt |> Spec_Rules.get
       
   329     val (rec_defs, nonrec_defs) =
       
   330       specs |> filter (curry (op =) Spec_Rules.Equational o fst)
       
   331             |> maps (snd o snd)
       
   332             |> filter_out (member Thm.eq_thm_prop risky_defs)
       
   333             |> List.partition (is_rec_def o prop_of)
       
   334     val spec_intros =
       
   335       specs |> filter (member (op =) [Spec_Rules.Inductive,
       
   336                                       Spec_Rules.Co_Inductive] o fst)
       
   337             |> maps (snd o snd)
       
   338   in
       
   339     Termtab.empty |> add Simp [atomize] snd simps
       
   340                   |> add Simp [] I rec_defs
       
   341                   |> add Def [] I nonrec_defs
       
   342 (* Add once it is used:
       
   343                   |> add Elim [] I elims
       
   344 *)
       
   345                   |> add Intro [] I intros
       
   346                   |> add Inductive [] I spec_intros
       
   347   end
       
   348 
       
   349 fun uniquify xs =
       
   350   Termtab.fold (cons o snd)
       
   351                (fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty) []
       
   352 
       
   353 fun struct_induct_rule_on th =
       
   354   case Logic.strip_horn (prop_of th) of
       
   355     (prems, @{const Trueprop}
       
   356             $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
       
   357     if not (is_TVar ind_T) andalso length prems > 1 andalso
       
   358        exists (exists_subterm (curry (op aconv) p)) prems andalso
       
   359        not (exists (exists_subterm (curry (op aconv) a)) prems) then
       
   360       SOME (p_name, ind_T)
       
   361     else
       
   362       NONE
       
   363   | _ => NONE
       
   364 
       
   365 fun instantiate_induct_rule ctxt concl_prop p_name ((name, stature), th) ind_x =
       
   366   let
       
   367     fun varify_noninducts (t as Free (s, T)) =
       
   368         if (s, T) = ind_x orelse can dest_funT T then t else Var ((s, 0), T)
       
   369       | varify_noninducts t = t
       
   370     val p_inst =
       
   371       concl_prop |> map_aterms varify_noninducts |> close_form
       
   372                  |> lambda (Free ind_x)
       
   373                  |> string_for_term ctxt
       
   374   in
       
   375     ((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]",
       
   376       stature), th |> read_instantiate ctxt [((p_name, 0), p_inst)])
       
   377   end
       
   378 
       
   379 fun type_match thy (T1, T2) =
       
   380   (Sign.typ_match thy (T2, T1) Vartab.empty; true)
       
   381   handle Type.TYPE_MATCH => false
       
   382 
       
   383 fun instantiate_if_induct_rule ctxt stmt stmt_xs (ax as (_, th)) =
       
   384   case struct_induct_rule_on th of
       
   385     SOME (p_name, ind_T) =>
       
   386     let val thy = Proof_Context.theory_of ctxt in
       
   387       stmt_xs |> filter (fn (_, T) => type_match thy (T, ind_T))
       
   388               |> map_filter (try (instantiate_induct_rule ctxt stmt p_name ax))
       
   389     end
       
   390   | NONE => [ax]
       
   391 
       
   392 fun external_frees t =
       
   393   [] |> Term.add_frees t |> filter_out (can Name.dest_internal o fst)
       
   394 
       
   395 fun maybe_instantiate_inducts ctxt hyp_ts concl_t =
       
   396   if Config.get ctxt instantiate_inducts then
       
   397     let
       
   398       val thy = Proof_Context.theory_of ctxt
       
   399       val ind_stmt =
       
   400         (hyp_ts |> filter_out (null o external_frees), concl_t)
       
   401         |> Logic.list_implies |> Object_Logic.atomize_term thy
       
   402       val ind_stmt_xs = external_frees ind_stmt
       
   403     in maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs) end
       
   404   else
       
   405     I
       
   406 
       
   407 fun maybe_filter_no_atps ctxt =
       
   408   not (Config.get ctxt ignore_no_atp) ? filter_out (No_ATPs.member ctxt o snd)
       
   409 
       
   410 fun nearly_all_facts ctxt ho_atp ({add, only, ...} : relevance_override)
       
   411                      chained_ths hyp_ts concl_t =
       
   412   if only andalso null add then
       
   413     []
       
   414   else
       
   415     let
       
   416       val reserved = reserved_isar_keyword_table ()
       
   417       val add_ths = Attrib.eval_thms ctxt add
       
   418       val css_table = clasimpset_rule_table_of ctxt
       
   419     in
       
   420       (if only then
       
   421          maps (map (fn ((name, stature), th) => ((K name, stature), th))
       
   422                o fact_from_ref ctxt reserved chained_ths css_table) add
       
   423        else
       
   424          all_facts ctxt ho_atp reserved false add_ths chained_ths css_table)
       
   425       |> maybe_instantiate_inducts ctxt hyp_ts concl_t
       
   426       |> not only ? maybe_filter_no_atps ctxt
       
   427       |> uniquify
       
   428     end
       
   429 
    26 
   430 val relevant_facts = iterative_relevant_facts
    27 val relevant_facts = iterative_relevant_facts
   431 
    28 
   432 end;
    29 end;