src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Mon, 23 Aug 2010 14:54:17 +0200
changeset 38652 e063be321438
parent 38644 25bbbaf7ce65
child 38679 2cfd0777580f
permissions -rw-r--r--
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later; it's a mistake to transform the elim rules too early because then we lose some info, e.g. "no_atp" attributes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 37995
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
36393
be73a2b2443b support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents: 36227
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
     4
*)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     5
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     6
signature SLEDGEHAMMER_FACT_FILTER =
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     7
sig
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
     8
  type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
     9
    {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    10
     del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    11
     only: bool}
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    12
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    13
  val trace : bool Unsynchronized.ref
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    14
  val chained_prefix : string
38617
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    15
  val name_thms_pair_from_ref :
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    16
    Proof.context -> thm list -> Facts.ref -> string * thm list
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
    17
  val relevant_facts :
37580
c2c1caff5dea got rid of "respect_no_atp" option, which even I don't use
blanchet
parents: 37578
diff changeset
    18
    bool -> real -> real -> bool -> int -> bool -> relevance_override
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
    19
    -> Proof.context * (thm list * 'a) -> term list -> term
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
    20
    -> (string * thm) list
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    21
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    22
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    23
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    24
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    25
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
    26
open Sledgehammer_Util
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
    27
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    28
val trace = Unsynchronized.ref false
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    29
fun trace_msg msg = if !trace then tracing (msg ()) else ()
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    30
37580
c2c1caff5dea got rid of "respect_no_atp" option, which even I don't use
blanchet
parents: 37578
diff changeset
    31
val respect_no_atp = true
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    32
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    33
type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    34
  {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    35
   del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    36
   only: bool}
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
    37
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    38
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    39
(* Used to label theorems chained into the goal. *)
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    40
val chained_prefix = sledgehammer_prefix ^ "chained_"
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    41
38617
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    42
fun name_thms_pair_from_ref ctxt chained_ths xref =
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    43
  let
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    44
    val ths = ProofContext.get_fact ctxt xref
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    45
    val name = Facts.string_of_ref xref
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    46
               |> forall (member Thm.eq_thm chained_ths) ths
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    47
                  ? prefix chained_prefix
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
    48
  in (name, ths) end
38617
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
    49
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    50
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    51
(***************************************************************)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    52
(* Relevance Filtering                                         *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    53
(***************************************************************)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    54
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    55
(*** constants with types ***)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    56
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    57
(*An abstraction of Isabelle types*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    58
datatype const_typ =  CTVar | CType of string * const_typ list
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    59
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    60
(*Is the second type an instance of the first one?*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    61
fun match_type (CType(con1,args1)) (CType(con2,args2)) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    62
      con1=con2 andalso match_types args1 args2
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    63
  | match_type CTVar _ = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    64
  | match_type _ CTVar = false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    65
and match_types [] [] = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    66
  | match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    67
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    68
(*Is there a unifiable constant?*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    69
fun uni_mem goal_const_tab (c, c_typ) =
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    70
  exists (match_types c_typ) (these (Symtab.lookup goal_const_tab c))
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    71
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    72
(*Maps a "real" type to a const_typ*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    73
fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    74
  | const_typ_of (TFree _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    75
  | const_typ_of (TVar _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    76
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    77
(*Pairs a constant with the list of its type instantiations (using const_typ)*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    78
fun const_with_typ thy (c,typ) =
38606
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
    79
  let val tvars = Sign.const_typargs thy (c,typ) in
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
    80
    (c, map const_typ_of tvars) end
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
    81
  handle TYPE _ => (c, [])   (*Variable (locale constant): monomorphic*)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    82
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    83
(*Add a const/type pair to the table, but a [] entry means a standard connective,
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    84
  which we ignore.*)
37502
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    85
fun add_const_type_to_table (c, ctyps) =
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    86
  Symtab.map_default (c, [ctyps])
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    87
                     (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    88
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
    89
val fresh_prefix = "Sledgehammer.FRESH."
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    90
val flip = Option.map not
38091
0508ff84036f work around atomization failures
blanchet
parents: 38090
diff changeset
    91
(* These are typically simplified away by "Meson.presimplify". *)
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
    92
val boring_consts = [@{const_name If}, @{const_name Let}]
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    93
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    94
fun get_consts_typs thy pos ts =
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    95
  let
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
    96
    (* We include free variables, as well as constants, to handle locales. For
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
    97
       each quantifiers that must necessarily be skolemized by the ATP, we
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
    98
       introduce a fresh constant to simulate the effect of Skolemization. *)
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    99
    fun do_term t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   100
      case t of
38606
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
   101
        Const x => add_const_type_to_table (const_with_typ thy x)
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
   102
      | Free (s, _) => add_const_type_to_table (s, [])
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   103
      | t1 $ t2 => do_term t1 #> do_term t2
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   104
      | Abs (_, _, t') => do_term t'
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   105
      | _ => I
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   106
    fun do_quantifier will_surely_be_skolemized body_t =
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   107
      do_formula pos body_t
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   108
      #> (if will_surely_be_skolemized then
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   109
            add_const_type_to_table (gensym fresh_prefix, [])
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   110
          else
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   111
            I)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   112
    and do_term_or_formula T =
38592
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   113
      if T = @{typ bool} orelse T = @{typ prop} then do_formula NONE
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   114
      else do_term
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   115
    and do_formula pos t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   116
      case t of
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   117
        Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   118
        do_quantifier (pos = SOME false) body_t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   119
      | @{const "==>"} $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   120
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   121
      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   122
        fold (do_term_or_formula T) [t1, t2]
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   123
      | @{const Trueprop} $ t1 => do_formula pos t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   124
      | @{const Not} $ t1 => do_formula (flip pos) t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   125
      | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   126
        do_quantifier (pos = SOME false) body_t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   127
      | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   128
        do_quantifier (pos = SOME true) body_t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   129
      | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   130
      | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   131
      | @{const "op -->"} $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   132
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   133
      | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   134
        fold (do_term_or_formula T) [t1, t2]
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   135
      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   136
        $ t1 $ t2 $ t3 =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   137
        do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   138
      | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   139
        do_quantifier (is_some pos) body_t
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   140
      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   141
        do_quantifier (pos = SOME false)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   142
                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   143
      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   144
        do_quantifier (pos = SOME true)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   145
                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   146
      | (t0 as Const (_, @{typ bool})) $ t1 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   147
        do_term t0 #> do_formula pos t1  (* theory constant *)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   148
      | _ => do_term t
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   149
  in
38091
0508ff84036f work around atomization failures
blanchet
parents: 38090
diff changeset
   150
    Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
   151
                 |> fold (do_formula pos) ts
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   152
  end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   153
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   154
(*Inserts a dummy "constant" referring to the theory name, so that relevance
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   155
  takes the given theory into account.*)
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   156
fun theory_const_prop_of theory_relevant th =
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   157
  if theory_relevant then
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   158
    let
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   159
      val name = Context.theory_name (theory_of_thm th)
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   160
      val t = Const (name ^ ". 1", @{typ bool})
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   161
    in t $ prop_of th end
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   162
  else
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   163
    prop_of th
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   164
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   165
(**** Constant / Type Frequencies ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   166
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   167
(*A two-dimensional symbol table counts frequencies of constants. It's keyed first by
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   168
  constant name and second by its list of type instantiations. For the latter, we need
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   169
  a linear ordering on type const_typ list.*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   170
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   171
local
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   172
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   173
fun cons_nr CTVar = 0
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   174
  | cons_nr (CType _) = 1;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   175
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   176
in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   177
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   178
fun const_typ_ord TU =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   179
  case TU of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   180
    (CType (a, Ts), CType (b, Us)) =>
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   181
      (case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   182
  | (T, U) => int_ord (cons_nr T, cons_nr U);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   183
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   184
end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   185
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31910
diff changeset
   186
structure CTtab = Table(type key = const_typ list val ord = dict_ord const_typ_ord);
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   187
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   188
fun count_axiom_consts theory_relevant thy (_, th) =
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   189
  let
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   190
    fun do_const (a, T) =
38606
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38595
diff changeset
   191
      let val (c, cts) = const_with_typ thy (a, T) in
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   192
        (* Two-dimensional table update. Constant maps to types maps to
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   193
           count. *)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   194
        CTtab.map_default (cts, 0) (Integer.add 1)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   195
        |> Symtab.map_default (c, CTtab.empty)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   196
      end
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   197
    fun do_term (Const x) = do_const x
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   198
      | do_term (Free x) = do_const x
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   199
      | do_term (t $ u) = do_term t #> do_term u
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   200
      | do_term (Abs (_, _, t)) = do_term t
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   201
      | do_term _ = I
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   202
  in th |> theory_const_prop_of theory_relevant |> do_term end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   203
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   204
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   205
(**** Actual Filtering Code ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   206
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   207
(*The frequency of a constant is the sum of those of all instances of its type.*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   208
fun const_frequency const_tab (c, cts) =
36185
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   209
  CTtab.fold (fn (cts', m) => match_types cts cts' ? Integer.add m)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   210
             (the (Symtab.lookup const_tab c)
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   211
              handle Option.Option => raise Fail ("Const: " ^ c)) 0
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   212
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   213
(* A surprising number of theorems contain only a few significant constants.
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   214
   These include all induction rules, and other general theorems. *)
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   215
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   216
(* "log" seems best in practice. A constant function of one ignores the constant
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   217
   frequencies. *)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   218
fun log_weight2 (x:real) = 1.0 + 2.0 / Math.ln (x + 1.0)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   219
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   220
(* Computes a constant's weight, as determined by its frequency. *)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   221
val ct_weight = log_weight2 o real oo const_frequency
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   222
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   223
(*Relevant constants are weighted according to frequency,
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   224
  but irrelevant constants are simply counted. Otherwise, Skolem functions,
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   225
  which are rare, would harm a formula's chances of being picked.*)
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   226
fun formula_weight const_tab gctyps consts_typs =
38101
34b75b71235d handle division by zero gracefully (used to raise Unordered later on)
blanchet
parents: 38095
diff changeset
   227
  let
34b75b71235d handle division by zero gracefully (used to raise Unordered later on)
blanchet
parents: 38095
diff changeset
   228
    val rel = filter (uni_mem gctyps) consts_typs
34b75b71235d handle division by zero gracefully (used to raise Unordered later on)
blanchet
parents: 38095
diff changeset
   229
    val rel_weight = fold (curry Real.+ o ct_weight const_tab) rel 0.0
34b75b71235d handle division by zero gracefully (used to raise Unordered later on)
blanchet
parents: 38095
diff changeset
   230
    val res = rel_weight / (rel_weight + real (length consts_typs - length rel))
34b75b71235d handle division by zero gracefully (used to raise Unordered later on)
blanchet
parents: 38095
diff changeset
   231
  in if Real.isFinite res then res else 0.0 end
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   232
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   233
(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   234
fun add_expand_pairs (x,ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   235
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   236
fun consts_typs_of_term thy t =
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   237
  Symtab.fold add_expand_pairs (get_consts_typs thy (SOME true) [t]) []
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   238
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   239
fun pair_consts_typs_axiom theory_relevant thy axiom =
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   240
  (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   241
                |> consts_typs_of_term thy)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   242
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   243
exception CONST_OR_FREE of unit
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   244
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   245
fun dest_Const_or_Free (Const x) = x
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   246
  | dest_Const_or_Free (Free x) = x
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   247
  | dest_Const_or_Free _ = raise CONST_OR_FREE ()
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   248
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   249
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   250
fun defines thy thm gctypes =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   251
    let val tm = prop_of thm
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   252
        fun defs lhs rhs =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   253
            let val (rator,args) = strip_comb lhs
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   254
                val ct = const_with_typ thy (dest_Const_or_Free rator)
33037
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 32994
diff changeset
   255
            in
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 32994
diff changeset
   256
              forall is_Var args andalso uni_mem gctypes ct andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   257
                subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   258
            end
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   259
            handle CONST_OR_FREE () => false
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   260
    in
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   261
        case tm of
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   262
          @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) =>
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   263
            defs lhs rhs
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   264
        | _ => false
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   265
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   266
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   267
type annotated_thm = (string * thm) * (string * const_typ list) list
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   268
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   269
(*For a reverse sort, putting the largest values first.*)
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   270
fun compare_pairs ((_, w1), (_, w2)) = Real.compare (w2, w1)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   271
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   272
(* Limit the number of new facts, to prevent runaway acceptance. *)
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   273
fun take_best max_new (newpairs : (annotated_thm * real) list) =
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   274
  let val nnew = length newpairs in
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   275
    if nnew <= max_new then
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   276
      (map #1 newpairs, [])
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   277
    else
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   278
      let
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   279
        val newpairs = sort compare_pairs newpairs
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   280
        val accepted = List.take (newpairs, max_new)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   281
      in
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   282
        trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   283
                       ", exceeds the limit of " ^ Int.toString max_new));
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   284
        trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   285
        trace_msg (fn () => "Actually passed: " ^
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   286
          space_implode ", " (map (fst o fst o fst) accepted));
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   287
        (map #1 accepted, map #1 (List.drop (newpairs, max_new)))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   288
      end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   289
  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   290
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   291
fun relevance_filter ctxt relevance_threshold relevance_convergence
38594
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   292
                     defs_relevant max_new theory_relevant
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   293
                     ({add, del, ...} : relevance_override) axioms goal_ts =
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   294
  if relevance_threshold > 1.0 then
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   295
    []
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   296
  else if relevance_threshold < 0.0 then
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   297
    axioms
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   298
  else
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   299
    let
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   300
      val thy = ProofContext.theory_of ctxt
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   301
      val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   302
                           Symtab.empty
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   303
      val goal_const_tab = get_consts_typs thy (SOME false) goal_ts
38090
fe51c098b0ab fiddle with the fudge factors, to get similar results as before
blanchet
parents: 38085
diff changeset
   304
      val relevance_threshold = 0.8 * relevance_threshold (* FIXME *)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   305
      val _ =
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   306
        trace_msg (fn () => "Initial constants: " ^
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   307
                            commas (goal_const_tab
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   308
                                    |> Symtab.dest
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   309
                                    |> filter (curry (op <>) [] o snd)
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   310
                                    |> map fst))
38594
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   311
      val add_thms = maps (ProofContext.get_fact ctxt) add
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   312
      val del_thms = maps (ProofContext.get_fact ctxt) del
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   313
      fun iter threshold rel_const_tab =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   314
        let
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   315
          fun relevant ([], rejects) [] =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   316
              (* Nothing was added this iteration: Add "add:" facts. *)
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   317
              if null add_thms then
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   318
                []
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   319
              else
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   320
                map_filter (fn (p as (name, th), _) =>
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   321
                               if member Thm.eq_thm add_thms th then SOME p
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   322
                               else NONE) rejects
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   323
            | relevant (newpairs, rejects) [] =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   324
              let
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   325
                val (newrels, more_rejects) = take_best max_new newpairs
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   326
                val new_consts = maps #2 newrels
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   327
                val rel_const_tab =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   328
                  rel_const_tab |> fold add_const_type_to_table new_consts
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   329
                val threshold =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   330
                  threshold + (1.0 - threshold) / relevance_convergence
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   331
              in
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   332
                trace_msg (fn () => "relevant this iteration: " ^
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   333
                                    Int.toString (length newrels));
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   334
                map #1 newrels @ iter threshold rel_const_tab
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   335
                    (more_rejects @ rejects)
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   336
              end
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   337
            | relevant (newrels, rejects)
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   338
                       ((ax as ((name, th), consts_typs)) :: axs) =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   339
              let
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   340
                val weight =
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   341
                  if member Thm.eq_thm del_thms th then 0.0
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   342
                  else formula_weight const_tab rel_const_tab consts_typs
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   343
              in
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   344
                if weight >= threshold orelse
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   345
                   (defs_relevant andalso defines thy th rel_const_tab) then
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   346
                  (trace_msg (fn () =>
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   347
                       name ^ " passes: " ^ Real.toString weight
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   348
                       (* ^ " consts: " ^ commas (map fst consts_typs) *));
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   349
                   relevant ((ax, weight) :: newrels, rejects) axs)
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   350
                else
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   351
                  relevant (newrels, ax :: rejects) axs
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   352
              end
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   353
          in
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   354
            trace_msg (fn () => "relevant_facts, current threshold: " ^
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   355
                                Real.toString threshold);
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   356
            relevant ([], [])
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   357
          end
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   358
      val relevant = iter relevance_threshold goal_const_tab
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   359
                          (map (pair_consts_typs_axiom theory_relevant thy)
af205f4fd0f3 tuning of relevance filter
blanchet
parents: 38593
diff changeset
   360
                               axioms)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   361
    in
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   362
      trace_msg (fn () => "Total relevant: " ^ Int.toString (length relevant));
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   363
      relevant
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   364
    end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   365
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   366
(***************************************************************)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   367
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   368
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   369
33022
c95102496490 Removal of the unused atpset concept, the atp attribute and some related code.
paulson
parents: 32994
diff changeset
   370
(*** retrieve lemmas and filter them ***)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   371
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   372
(*Reject theorems with names like "List.filter.filter_list_def" or
21690
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   373
  "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   374
fun is_package_def a =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30291
diff changeset
   375
  let val names = Long_Name.explode a
21690
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   376
  in
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   377
     length names > 2 andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   378
     not (hd names = "local") andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   379
     String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   380
  end;
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   381
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   382
fun make_fact_table xs =
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   383
  fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   384
fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   385
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   386
(* FIXME: put other record thms here, or declare as "no_atp" *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   387
val multi_base_blacklist =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   388
  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   389
   "split_asm", "cases", "ext_cases"]
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   390
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   391
val max_lambda_nesting = 3
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   392
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   393
fun term_has_too_many_lambdas max (t1 $ t2) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   394
    exists (term_has_too_many_lambdas max) [t1, t2]
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   395
  | term_has_too_many_lambdas max (Abs (_, _, t)) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   396
    max = 0 orelse term_has_too_many_lambdas (max - 1) t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   397
  | term_has_too_many_lambdas _ _ = false
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   398
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   399
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   400
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   401
(* Don't count nested lambdas at the level of formulas, since they are
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   402
   quantifiers. *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   403
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   404
    formula_has_too_many_lambdas (T :: Ts) t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   405
  | formula_has_too_many_lambdas Ts t =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   406
    if is_formula_type (fastype_of1 (Ts, t)) then
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   407
      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   408
    else
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   409
      term_has_too_many_lambdas max_lambda_nesting t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   410
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   411
(* The max apply depth of any "metis" call in "Metis_Examples" (on 31-10-2007)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   412
   was 11. *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   413
val max_apply_depth = 15
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   414
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   415
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   416
  | apply_depth (Abs (_, _, t)) = apply_depth t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   417
  | apply_depth _ = 0
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   418
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   419
fun is_formula_too_complex t =
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   420
  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   421
37543
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   422
val exists_sledgehammer_const =
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   423
  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   424
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   425
fun is_strange_theorem th =
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   426
  case head_of (concl_of th) of
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   427
      Const (a, _) => (a <> @{const_name Trueprop} andalso
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   428
                       a <> @{const_name "=="})
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   429
    | _ => false
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   430
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   431
val type_has_top_sort =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   432
  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   433
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   434
(**** Predicates to detect unwanted facts (prolific or likely to cause
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   435
      unsoundness) ****)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   436
38289
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   437
(* Too general means, positive equality literal with a variable X as one
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   438
   operand, when X does not occur properly in the other operand. This rules out
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   439
   clearly inconsistent facts such as X = a | X = b, though it by no means
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   440
   guarantees soundness. *)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   441
38289
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   442
(* Unwanted equalities are those between a (bound or schematic) variable that
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   443
   does not properly occur in the second operand. *)
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   444
val is_exhaustive_finite =
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   445
  let
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   446
    fun is_bad_equal (Var z) t =
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   447
        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   448
      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   449
      | is_bad_equal _ _ = false
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   450
    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   451
    fun do_formula pos t =
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   452
      case (pos, t) of
38615
4e1d828ee514 improve "x = A | x = B | x = C"-style axiom detection
blanchet
parents: 38611
diff changeset
   453
        (_, @{const Trueprop} $ t1) => do_formula pos t1
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   454
      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   455
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   456
      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   457
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   458
      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   459
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   460
      | (_, @{const "==>"} $ t1 $ t2) =>
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   461
        do_formula (not pos) t1 andalso
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   462
        (t2 = @{prop False} orelse do_formula pos t2)
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   463
      | (_, @{const "op -->"} $ t1 $ t2) =>
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   464
        do_formula (not pos) t1 andalso
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   465
        (t2 = @{const False} orelse do_formula pos t2)
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   466
      | (_, @{const Not} $ t1) => do_formula (not pos) t1
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   467
      | (true, @{const "op |"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   468
      | (false, @{const "op &"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   469
      | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   470
      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   471
      | _ => false
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   472
  in do_formula true end
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   473
38592
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   474
fun has_bound_or_var_of_type tycons =
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   475
  exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   476
                   | Abs (_, Type (s, _), _) => member (op =) tycons s
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   477
                   | _ => false)
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   478
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   479
(* Facts are forbidden to contain variables of these types. The typical reason
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   480
   is that they lead to unsoundness. Note that "unit" satisfies numerous
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   481
   equations like "?x = ()". The resulting clauses will have no type constraint,
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   482
   yielding false proofs. Even "bool" leads to many unsound proofs, though only
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   483
   for higher-order problems. *)
38592
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   484
val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   485
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   486
(* Facts containing variables of type "unit" or "bool" or of the form
38290
581a402a80f0 prevent ATP thread for staying around for 1 minute if an exception occurred earlier;
blanchet
parents: 38289
diff changeset
   487
   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
581a402a80f0 prevent ATP thread for staying around for 1 minute if an exception occurred earlier;
blanchet
parents: 38289
diff changeset
   488
   are omitted. *)
38593
85aef7587cf1 minor cleanup
blanchet
parents: 38592
diff changeset
   489
fun is_dangerous_term full_types t =
38609
6220e5ab32f7 more fiddling with Sledgehammer's "add:" option
blanchet
parents: 38607
diff changeset
   490
  not full_types andalso
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   491
  ((has_bound_or_var_of_type dangerous_types t andalso
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   492
    has_bound_or_var_of_type dangerous_types (transform_elim_term t))
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   493
   orelse is_exhaustive_finite t)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   494
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   495
fun is_theorem_bad_for_atps full_types thm =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   496
  let val t = prop_of thm in
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   497
    is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   498
    is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   499
    is_strange_theorem thm
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   500
  end
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   501
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   502
fun all_name_thms_pairs ctxt full_types add_thms chained_ths =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   503
  let
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   504
    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   505
    val local_facts = ProofContext.facts_of ctxt
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   506
    val named_locals = local_facts |> Facts.dest_static []
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   507
    val unnamed_locals =
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   508
      local_facts |> Facts.props
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   509
      |> filter_out (fn th => exists (fn (_, ths) => member Thm.eq_thm ths th)
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   510
                                     named_locals)
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   511
      |> map (pair "" o single)
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   512
    val full_space =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   513
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   514
    fun valid_facts facts pairs =
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   515
      (pairs, []) |-> fold (fn (name, ths0) =>
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   516
        if name <> "" andalso
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   517
           forall (not o member Thm.eq_thm add_thms) ths0 andalso
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   518
           (Facts.is_concealed facts name orelse
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   519
            (respect_no_atp andalso is_package_def name) orelse
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   520
            member (op =) multi_base_blacklist (Long_Name.base_name name) orelse
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   521
            String.isSuffix "_def_raw" (* FIXME: crude hack *) name) then
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   522
          I
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   523
        else
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   524
          let
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   525
            fun check_thms a =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   526
              (case try (ProofContext.get_thms ctxt) a of
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   527
                NONE => false
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   528
              | SOME ths1 => Thm.eq_thms (ths0, ths1))
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   529
            val name1 = Facts.extern facts name;
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   530
            val name2 = Name_Space.extern full_space name;
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   531
            val ths =
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   532
              ths0 |> filter ((not o is_theorem_bad_for_atps full_types) orf
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   533
                              member Thm.eq_thm add_thms)
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   534
            val name' =
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   535
              case find_first check_thms [name1, name2, name] of
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   536
                SOME name' => name'
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   537
              | NONE =>
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   538
                ths |> map (fn th =>
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   539
                               "`" ^ Print_Mode.setmp [Print_Mode.input]
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   540
                                         (Syntax.string_of_term ctxt)
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   541
                                         (prop_of th) ^ "`")
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   542
                    |> space_implode " "
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   543
          in
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
   544
            cons (name' |> forall (member Thm.eq_thm chained_ths) ths0
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   545
                           ? prefix chained_prefix, ths)
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   546
          end)
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   547
  in
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   548
    valid_facts local_facts (unnamed_locals @ named_locals) @
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   549
    valid_facts global_facts (Facts.dest_static [] global_facts)
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   550
  end
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   551
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   552
fun multi_name a th (n, pairs) =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   553
  (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   554
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   555
fun add_names (_, []) pairs = pairs
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   556
  | add_names (a, [th]) pairs = (a, th) :: pairs
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   557
  | add_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs))
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   558
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   559
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   560
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   561
(* The single-name theorems go after the multiple-name ones, so that single
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   562
   names are preferred when both are available. *)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   563
fun name_thm_pairs ctxt respect_no_atp name_thms_pairs =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   564
  let
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   565
    val (mults, singles) = List.partition is_multi name_thms_pairs
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   566
    val ps = [] |> fold add_names singles |> fold add_names mults
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   567
  in ps |> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) end;
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   568
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   569
fun is_named ("", th) =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   570
    (warning ("No name for theorem " ^
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   571
              Display.string_of_thm_without_context th); false)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   572
  | is_named _ = true
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   573
fun checked_name_thm_pairs respect_no_atp ctxt =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   574
  name_thm_pairs ctxt respect_no_atp
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   575
  #> tap (fn ps => trace_msg
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   576
                        (fn () => ("Considering " ^ Int.toString (length ps) ^
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   577
                                   " theorems")))
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   578
  #> filter is_named
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   579
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   580
(***************************************************************)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   581
(* ATP invocation methods setup                                *)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   582
(***************************************************************)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   583
37580
c2c1caff5dea got rid of "respect_no_atp" option, which even I don't use
blanchet
parents: 37578
diff changeset
   584
fun relevant_facts full_types relevance_threshold relevance_convergence
c2c1caff5dea got rid of "respect_no_atp" option, which even I don't use
blanchet
parents: 37578
diff changeset
   585
                   defs_relevant max_new theory_relevant
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   586
                   (relevance_override as {add, del, only})
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
   587
                   (ctxt, (chained_ths, _)) hyp_ts concl_t =
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   588
  let
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   589
    val add_thms = maps (ProofContext.get_fact ctxt) add
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   590
    val axioms =
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   591
      checked_name_thm_pairs (respect_no_atp andalso not only) ctxt
38617
f7b32911340b unbreak "only" option of Sledgehammer
blanchet
parents: 38615
diff changeset
   592
          (if only then map (name_thms_pair_from_ref ctxt chained_ths) add
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   593
           else all_name_thms_pairs ctxt full_types add_thms chained_ths)
38595
bbb0982656eb make sure that "add:" doesn't influence the relevance filter too much
blanchet
parents: 38594
diff changeset
   594
      |> make_unique
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   595
  in
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   596
    relevance_filter ctxt relevance_threshold relevance_convergence
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   597
                     defs_relevant max_new theory_relevant relevance_override
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   598
                     axioms (concl_t :: hyp_ts)
38595
bbb0982656eb make sure that "add:" doesn't influence the relevance filter too much
blanchet
parents: 38594
diff changeset
   599
    |> sort_wrt fst
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   600
  end
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   601
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   602
end;