src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Fri, 25 Jun 2010 12:08:48 +0200
changeset 37551 2dc53a9f69c9
parent 37543 2e733b0a963c
child 37552 6034aac65143
permissions -rw-r--r--
improve the natural formula relevance filter code, so that it behaves more like the CNF one
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
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory, 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
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
     8
  type cnf_thm = Sledgehammer_Fact_Preprocessor.cnf_thm
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
     9
  type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    10
  type arity_clause = Sledgehammer_FOL_Clause.arity_clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    11
  type hol_clause = Sledgehammer_HOL_Clause.hol_clause
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    12
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    13
  type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    14
    {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    15
     del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    16
     only: bool}
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    17
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    18
  val use_natural_form : bool Unsynchronized.ref
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
    19
  val name_thms_pair_from_ref :
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
    20
    Proof.context -> thm list -> Facts.ref -> string * thm list
22989
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    21
  val tvar_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    22
  val tfree_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    23
  val type_consts_of_terms : theory -> term list -> string list
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
    24
  val is_quasi_fol_theorem : theory -> thm -> bool
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
    25
  val relevant_facts :
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
    26
    bool -> bool -> real -> real -> bool -> int -> bool -> relevance_override
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    27
    -> Proof.context * (thm list * 'a) -> thm list -> cnf_thm list
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    28
  val prepare_clauses :
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    29
    bool -> thm list -> cnf_thm list -> cnf_thm list -> theory
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    30
    -> string vector
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    31
       * (hol_clause list * hol_clause list * hol_clause list
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    32
          * hol_clause list * classrel_clause list * arity_clause list)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    33
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    34
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    35
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    36
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    37
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    38
open Sledgehammer_FOL_Clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    39
open Sledgehammer_Fact_Preprocessor
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    40
open Sledgehammer_HOL_Clause
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    41
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    42
(* Experimental feature: Filter theorems in natural form or in CNF? *)
37506
32a1ee39c49b missing "Unsynchronized" + make exception take a unit
blanchet
parents: 37505
diff changeset
    43
val use_natural_form = Unsynchronized.ref false
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    44
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    45
type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    46
  {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    47
   del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    48
   only: bool}
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
    49
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    50
(***************************************************************)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    51
(* Relevance Filtering                                         *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    52
(***************************************************************)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    53
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    54
fun strip_Trueprop (@{const Trueprop} $ t) = t
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
    55
  | strip_Trueprop t = t;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    56
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    57
(*** constants with types ***)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    58
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    59
(*An abstraction of Isabelle types*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    60
datatype const_typ =  CTVar | CType of string * const_typ list
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    61
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    62
(*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
    63
fun match_type (CType(con1,args1)) (CType(con2,args2)) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    64
      con1=con2 andalso match_types args1 args2
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    65
  | match_type CTVar _ = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    66
  | match_type _ CTVar = false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    67
and match_types [] [] = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    68
  | 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
    69
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    70
(*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
    71
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
    72
  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
    73
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    74
(*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
    75
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
    76
  | const_typ_of (TFree _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    77
  | const_typ_of (TVar _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    78
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    79
(*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
    80
fun const_with_typ thy (c,typ) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    81
    let val tvars = Sign.const_typargs thy (c,typ)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    82
    in (c, map const_typ_of tvars) end
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    83
    handle TYPE _ => (c,[]);   (*Variable (locale constant): monomorphic*)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    84
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    85
(*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
    86
  which we ignore.*)
37502
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    87
fun add_const_type_to_table (c, ctyps) =
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    88
  Symtab.map_default (c, [ctyps])
a8f7b25d5478 canonical argument order
blanchet
parents: 37501
diff changeset
    89
                     (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    90
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
    91
val fresh_prefix = "Sledgehammer.Fresh."
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    92
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    93
val flip = Option.map not
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    94
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
    95
val boring_natural_consts = [@{const_name If}]
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    96
(* Including equality in this list might be expected to stop rules like
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    97
   subset_antisym from being chosen, but for some reason filtering works better
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    98
   with them listed. The logical signs All, Ex, &, and --> are omitted for CNF
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
    99
   because any remaining occurrences must be within comprehensions. *)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   100
val boring_cnf_consts =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   101
  [@{const_name Trueprop}, @{const_name "==>"}, @{const_name all},
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   102
   @{const_name "=="}, @{const_name "op |"}, @{const_name Not},
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   103
   @{const_name "op ="}]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   104
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   105
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
   106
  let
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   107
    (* Free variables are included, as well as constants, to handle locales.
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   108
       "skolem_id" is included to increase the complexity of theorems containing
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   109
       Skolem functions. In non-CNF form, "Ex" is included but each occurrence
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   110
       is considered fresh, to simulate the effect of Skolemization. *)
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   111
    fun do_term t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   112
      case t of
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   113
        Const x => add_const_type_to_table (const_with_typ thy x)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   114
      | Free x => add_const_type_to_table (const_with_typ thy x)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   115
      | (t as Const (@{const_name skolem_id}, _)) $ _ => do_term t
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   116
      | t1 $ t2 => do_term t1 #> do_term t2
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   117
      | Abs (_, _, t) =>
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   118
        (* Abstractions lead to combinators, so we add a penalty for them. *)
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   119
        add_const_type_to_table (gensym fresh_prefix, [])
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   120
        #> do_term t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   121
      | _ => I
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   122
    fun do_quantifier sweet_pos pos body_t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   123
      do_formula pos body_t
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   124
      #> (if pos = SOME sweet_pos then I
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   125
          else add_const_type_to_table (gensym fresh_prefix, []))
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   126
    and do_equality T t1 t2 =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   127
      fold (if T = @{typ bool} orelse T = @{typ prop} then do_formula NONE
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   128
            else do_term) [t1, t2]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   129
    and do_formula pos t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   130
      case t of
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   131
        Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   132
        do_quantifier false pos body_t
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   133
      | @{const "==>"} $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   134
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   135
      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   136
        do_equality T t1 t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   137
      | @{const Trueprop} $ t1 => do_formula pos t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   138
      | @{const Not} $ t1 => do_formula (flip pos) t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   139
      | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   140
        do_quantifier false pos body_t
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   141
      | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   142
        do_quantifier true pos body_t
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   143
      | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   144
      | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   145
      | @{const "op -->"} $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   146
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   147
      | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   148
        do_equality T t1 t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   149
      | (t0 as Const (_, @{typ bool})) $ t1 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   150
        do_term t0 #> do_formula pos t1  (* theory constant *)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   151
      | _ => do_term t
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   152
  in
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   153
    Symtab.empty
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   154
    |> (if !use_natural_form then
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   155
          fold (Symtab.update o rpair []) boring_natural_consts
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   156
          #> fold (do_formula pos) ts
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   157
        else
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   158
          fold (Symtab.update o rpair []) boring_cnf_consts
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   159
          #> fold do_term ts)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   160
  end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   161
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   162
(*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
   163
  takes the given theory into account.*)
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   164
fun 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
   165
  if theory_relevant then
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   166
    let
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   167
      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
   168
      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
   169
    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
   170
  else
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   171
    prop_of th
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   172
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   173
fun appropriate_prop_of theory_relevant (cnf_thm, (_, orig_thm)) =
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   174
  (if !use_natural_form then orig_thm else cnf_thm)
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   175
  |> const_prop_of theory_relevant
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   176
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   177
(**** Constant / Type Frequencies ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   178
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   179
(*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
   180
  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
   181
  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
   182
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   183
local
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   184
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   185
fun cons_nr CTVar = 0
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   186
  | cons_nr (CType _) = 1;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   187
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   188
in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   189
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   190
fun const_typ_ord TU =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   191
  case TU of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   192
    (CType (a, Ts), CType (b, Us)) =>
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   193
      (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
   194
  | (T, U) => int_ord (cons_nr T, cons_nr U);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   195
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   196
end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   197
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31910
diff changeset
   198
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
   199
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   200
fun count_axiom_consts theory_relevant thy axiom =
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   201
  let
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   202
    fun do_const (a, T) =
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   203
      let val (c, cts) = const_with_typ thy (a,T) in
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   204
        (* 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
   205
           count. *)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   206
        CTtab.map_default (cts, 0) (Integer.add 1)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   207
        |> Symtab.map_default (c, CTtab.empty)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   208
      end
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   209
    fun do_term (Const x) = do_const x
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   210
      | do_term (Free x) = do_const x
37515
ef3742657bc6 fix bug with "skolem_id" + sort facts for increased readability
blanchet
parents: 37509
diff changeset
   211
      | do_term (Const (x as (@{const_name skolem_id}, _)) $ _) = do_const x
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   212
      | 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
   213
      | do_term (Abs (_, _, t)) = do_term t
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   214
      | do_term _ = I
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   215
  in axiom |> appropriate_prop_of theory_relevant |> do_term end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   216
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   217
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   218
(**** Actual Filtering Code ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   219
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   220
(*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
   221
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
   222
  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
   223
             (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
   224
              handle Option.Option => raise Fail ("Const: " ^ c)) 0
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   225
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   226
(*A surprising number of theorems contain only a few significant constants.
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   227
  These include all induction rules, and other general theorems. Filtering
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   228
  theorems in clause form reveals these complexities in the form of Skolem
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   229
  functions. If we were instead to filter theorems in their natural form,
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   230
  some other method of measuring theorem complexity would become necessary.*)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   231
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   232
(* "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
   233
   frequencies. *)
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   234
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
   235
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   236
(* 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
   237
val ct_weight = log_weight2 o real oo const_frequency
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
(*Relevant constants are weighted according to frequency,
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   240
  but irrelevant constants are simply counted. Otherwise, Skolem functions,
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   241
  which are rare, would harm a clause's chances of being picked.*)
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   242
fun clause_weight const_tab gctyps consts_typs =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   243
    let val rel = filter (uni_mem gctyps) consts_typs
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   244
        val rel_weight = fold (curry Real.+ o ct_weight const_tab) rel 0.0
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   245
    in
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   246
        rel_weight / (rel_weight + real (length consts_typs - length rel))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   247
    end;
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   248
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   249
(*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
   250
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
   251
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   252
fun consts_typs_of_term thy t =
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   253
  Symtab.fold add_expand_pairs (get_consts_typs thy (SOME false) [t]) []
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   254
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   255
fun pair_consts_typs_axiom theory_relevant thy axiom =
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   256
  (axiom, axiom |> appropriate_prop_of theory_relevant
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   257
                |> consts_typs_of_term thy)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   258
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   259
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
   260
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   261
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
   262
  | 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
   263
  | dest_Const_or_Free _ = raise CONST_OR_FREE ()
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   264
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   265
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   266
fun defines thy thm gctypes =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   267
    let val tm = prop_of thm
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   268
        fun defs lhs rhs =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   269
            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
   270
                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
   271
            in
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 32994
diff changeset
   272
              forall is_Var args andalso uni_mem gctypes ct andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   273
                subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   274
            end
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   275
            handle CONST_OR_FREE () => false
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   276
    in
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   277
        case tm of
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   278
          @{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
   279
            defs lhs rhs
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   280
        | _ => false
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   281
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   282
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   283
type annotated_clause = cnf_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
   284
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   285
(*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
   286
fun compare_pairs ((_, w1), (_, w2)) = Real.compare (w2, w1)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   288
(*Limit the number of new clauses, to prevent runaway acceptance.*)
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   289
fun take_best max_new (newpairs : (annotated_clause * real) list) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   290
  let val nnew = length newpairs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   291
  in
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   292
    if nnew <= max_new then (map #1 newpairs, [])
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   293
    else
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   294
      let val cls = sort compare_pairs newpairs
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   295
          val accepted = List.take (cls, max_new)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   296
      in
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   297
        trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   298
                       ", exceeds the limit of " ^ Int.toString (max_new)));
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   299
        trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   300
        trace_msg (fn () => "Actually passed: " ^
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   301
          space_implode ", " (map (fn (((_,((name,_), _)),_),_) => name) accepted));
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   302
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   303
        (map #1 accepted, map #1 (List.drop (cls, max_new)))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   304
      end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   305
  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   306
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   307
fun relevant_clauses ctxt relevance_convergence defs_relevant max_new
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   308
                     ({add, del, ...} : relevance_override) const_tab =
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   309
  let
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   310
    val thy = ProofContext.theory_of ctxt
37501
b5440c78ed3f leverage new data structure for handling "add:" and "del:"
blanchet
parents: 37500
diff changeset
   311
    val add_thms = maps (ProofContext.get_fact ctxt) add
b5440c78ed3f leverage new data structure for handling "add:" and "del:"
blanchet
parents: 37500
diff changeset
   312
    val del_thms = maps (ProofContext.get_fact ctxt) del
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   313
    fun iter threshold rel_const_tab =
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   314
      let
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   315
        fun relevant ([], _) [] = []  (* Nothing added this iteration *)
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   316
          | relevant (newpairs, rejects) [] =
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   317
            let
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   318
              val (newrels, more_rejects) = take_best max_new newpairs
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   319
              val new_consts = maps #2 newrels
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   320
              val rel_const_tab =
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   321
                rel_const_tab |> fold add_const_type_to_table new_consts
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   322
              val threshold =
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   323
                threshold + (1.0 - threshold) / relevance_convergence
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   324
            in
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   325
              trace_msg (fn () => "relevant this iteration: " ^
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   326
                                  Int.toString (length newrels));
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   327
              map #1 newrels @ iter threshold rel_const_tab
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   328
                  (more_rejects @ rejects)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   329
            end
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   330
          | relevant (newrels, rejects)
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   331
                     ((ax as (clsthm as (_, ((name, n), orig_th)),
37501
b5440c78ed3f leverage new data structure for handling "add:" and "del:"
blanchet
parents: 37500
diff changeset
   332
                              consts_typs)) :: axs) =
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   333
            let
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   334
              val weight =
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   335
                if member Thm.eq_thm add_thms orig_th then 1.0
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   336
                else if member Thm.eq_thm del_thms orig_th then 0.0
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   337
                else clause_weight const_tab rel_const_tab consts_typs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   338
            in
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   339
              if weight >= threshold orelse
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   340
                 (defs_relevant andalso
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   341
                  defines thy (#1 clsthm) rel_const_tab) then
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   342
                (trace_msg (fn () =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   343
                     name ^ " clause " ^ Int.toString n ^
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   344
                     " passes: " ^ Real.toString weight
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   345
                     (* ^ " consts: " ^ commas (map fst consts_typs) *));
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   346
                 relevant ((ax, weight) :: newrels, rejects) axs)
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   347
              else
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   348
                relevant (newrels, ax :: rejects) axs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   349
            end
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   350
        in
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   351
          trace_msg (fn () => "relevant_clauses, current threshold: " ^
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   352
                              Real.toString threshold);
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   353
          relevant ([], [])
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   354
        end
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   355
  in iter end
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   356
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   357
fun relevance_filter ctxt relevance_threshold relevance_convergence
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   358
                     defs_relevant max_new theory_relevant relevance_override
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   359
                     thy (axioms : cnf_thm list) goals =
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   360
  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
   361
    []
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   362
  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
   363
    axioms
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   364
  else
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   365
    let
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   366
      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
   367
                           Symtab.empty
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   368
      val goal_const_tab = get_consts_typs thy (SOME true) goals
37551
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   369
      val relevance_threshold =
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   370
        if !use_natural_form then 0.9 * relevance_threshold (* experimental *)
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   371
        else relevance_threshold
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   372
      val _ =
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   373
        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
   374
                            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
   375
                                    |> Symtab.dest
2dc53a9f69c9 improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet
parents: 37543
diff changeset
   376
                                    |> 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
   377
                                    |> map fst))
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   378
      val relevant =
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   379
        relevant_clauses ctxt relevance_convergence defs_relevant max_new
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   380
                         relevance_override const_tab relevance_threshold
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   381
                         goal_const_tab
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   382
                         (map (pair_consts_typs_axiom theory_relevant thy)
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   383
                              axioms)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   384
    in
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   385
      trace_msg (fn () => "Total relevant: " ^ Int.toString (length relevant));
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   386
      relevant
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   387
    end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   388
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   389
(***************************************************************)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   390
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   391
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   392
33022
c95102496490 Removal of the unused atpset concept, the atp attribute and some related code.
paulson
parents: 32994
diff changeset
   393
(*** retrieve lemmas and filter them ***)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   394
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   395
(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   396
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   397
fun setinsert (x,s) = Symtab.update (x,()) s;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   398
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   399
(*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
   400
  "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
   401
fun is_package_def a =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30291
diff changeset
   402
  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
   403
  in
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   404
     length names > 2 andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   405
     not (hd names = "local") andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   406
     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
   407
  end;
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   408
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   409
fun mk_clause_table xs =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   410
  fold (Termtab.update o `(prop_of o fst)) xs Termtab.empty
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   411
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   412
fun make_unique xs =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   413
  Termtab.fold (cons o snd) (mk_clause_table xs) []
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   414
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   415
(* Remove existing axiom clauses from the conjecture clauses, as this can
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   416
   dramatically boost an ATP's performance (for some reason). *)
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   417
fun subtract_cls ax_clauses =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   418
  filter_out (Termtab.defined (mk_clause_table ax_clauses) o prop_of)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   419
37543
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   420
val exists_sledgehammer_const =
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   421
  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s) o prop_of
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   422
37345
4402a2bfa204 make Sledgehammer's "add:" and "del:" syntax work better in the presence of aliases;
blanchet
parents: 37344
diff changeset
   423
fun all_name_thms_pairs respect_no_atp ctxt chained_ths =
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   424
  let
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   425
    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
26278
f0c6839df608 replaced obsolete FactIndex.T by Facts.T;
wenzelm
parents: 25761
diff changeset
   426
    val local_facts = ProofContext.facts_of ctxt;
33641
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   427
    val full_space =
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   428
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   429
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   430
    fun valid_facts facts =
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   431
      (facts, []) |-> Facts.fold_static (fn (name, ths0) =>
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   432
        if Facts.is_concealed facts name orelse
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   433
           (respect_no_atp andalso is_package_def name) orelse
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   434
           member (op =) multi_base_blacklist (Long_Name.base_name name) then
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   435
          I
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   436
        else
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   437
          let
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   438
            fun check_thms a =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   439
              (case try (ProofContext.get_thms ctxt) a of
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   440
                NONE => false
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   441
              | SOME ths1 => Thm.eq_thms (ths0, ths1));
33641
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   442
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   443
            val name1 = Facts.extern facts name;
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   444
            val name2 = Name_Space.extern full_space name;
37543
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   445
            val ths = filter_out (is_theorem_bad_for_atps orf
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   446
                                  exists_sledgehammer_const) ths0
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   447
          in
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   448
            case find_first check_thms [name1, name2, name] of
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   449
              NONE => I
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   450
            | SOME name' =>
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   451
              cons (name' |> forall (member Thm.eq_thm chained_ths) ths
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   452
                             ? prefix chained_prefix, ths)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   453
          end)
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   454
  in valid_facts global_facts @ valid_facts local_facts end;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   455
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   456
fun multi_name a th (n, pairs) =
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   457
  (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   458
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   459
fun add_names (_, []) pairs = pairs
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   460
  | add_names (a, [th]) pairs = (a, th) :: pairs
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   461
  | add_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs))
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   462
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   463
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   464
36550
f8da913b6c3a make sure short theorem names are preferred to composite ones in Sledgehammer;
blanchet
parents: 36473
diff changeset
   465
(* The single-name theorems go after the multiple-name ones, so that single
f8da913b6c3a make sure short theorem names are preferred to composite ones in Sledgehammer;
blanchet
parents: 36473
diff changeset
   466
   names are preferred when both are available. *)
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   467
fun name_thm_pairs respect_no_atp ctxt name_thms_pairs =
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   468
  let
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   469
    val (mults, singles) = List.partition is_multi name_thms_pairs
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   470
    val ps = [] |> fold add_names singles |> fold add_names mults
36060
4d27652ffb40 reintroduce efficient set structure to collect "no_atp" theorems
blanchet
parents: 36059
diff changeset
   471
  in ps |> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) end;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   472
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   473
fun is_named ("", th) =
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   474
    (warning ("No name for theorem " ^
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   475
              Display.string_of_thm_without_context th); false)
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   476
  | is_named _ = true
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   477
fun checked_name_thm_pairs respect_no_atp ctxt =
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   478
  name_thm_pairs respect_no_atp ctxt
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   479
  #> tap (fn ps => trace_msg
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   480
                        (fn () => ("Considering " ^ Int.toString (length ps) ^
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   481
                                   " theorems")))
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   482
  #> filter is_named
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   483
37344
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   484
fun name_thms_pair_from_ref ctxt chained_ths xref =
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   485
  let
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   486
    val ths = ProofContext.get_fact ctxt xref
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   487
    val name = Facts.string_of_ref xref
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   488
               |> forall (member Thm.eq_thm chained_ths) ths
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   489
                  ? prefix chained_prefix
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   490
  in (name, ths) end
40f379944c1e totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
blanchet
parents: 37171
diff changeset
   491
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   492
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   493
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   494
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   495
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   496
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32866
diff changeset
   497
fun add_classes (sorts, cset) = List.foldl setinsert cset (flat sorts);
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   498
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   499
(*Remove this trivial type class*)
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   500
fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   501
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   502
fun tvar_classes_of_terms ts =
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29267
diff changeset
   503
  let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   504
  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   505
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   506
fun tfree_classes_of_terms ts =
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29267
diff changeset
   507
  let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   508
  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   509
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   510
(*fold type constructors*)
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   511
fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   512
  | fold_type_consts _ _ x = x;
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   513
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   514
(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   515
fun add_type_consts_in_term thy =
37504
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   516
  let
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   517
    val const_typargs = Sign.const_typargs thy
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   518
    fun aux (Const cT) = fold (fold_type_consts setinsert) (const_typargs cT)
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   519
      | aux (Abs (_, _, u)) = aux u
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   520
      | aux (Const (@{const_name skolem_id}, _) $ _) = I
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   521
      | aux (t $ u) = aux t #> aux u
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   522
      | aux _ = I
4308d2bbbca8 more cosmetics
blanchet
parents: 37503
diff changeset
   523
  in aux end
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   524
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   525
fun type_consts_of_terms thy ts =
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   526
  Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   527
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   528
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   529
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   530
(* ATP invocation methods setup                                *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   531
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   532
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   533
fun is_quasi_fol_theorem thy =
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   534
  Meson.is_fol_term thy o snd o conceal_skolem_somes ~1 [] o prop_of
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   535
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
   536
(**** Predicates to detect unwanted clauses (prolific or likely to cause
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
   537
      unsoundness) ****)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   538
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   539
(** Too general means, positive equality literal with a variable X as one operand,
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   540
  when X does not occur properly in the other operand. This rules out clearly
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   541
  inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   542
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   543
fun var_occurs_in_term ix =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   544
  let
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   545
    fun aux (Var (jx, _)) = (ix = jx)
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   546
      | aux (t1 $ t2) = aux t1 orelse aux t2
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   547
      | aux (Abs (_, _, t)) = aux t
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   548
      | aux _ = false
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   549
  in aux end
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   550
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   551
fun is_record_type T = not (null (Record.dest_recTs T))
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   552
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   553
(*Unwanted equalities include
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   554
  (1) those between a variable that does not properly occur in the second operand,
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   555
  (2) those between a variable and a record, since these seem to be prolific "cases" thms
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   556
*)
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   557
fun too_general_eqterms (Var (ix,T), t) =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   558
    not (var_occurs_in_term ix t) orelse is_record_type T
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   559
  | too_general_eqterms _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   560
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   561
fun too_general_equality (Const (@{const_name "op ="}, _) $ x $ y) =
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   562
      too_general_eqterms (x,y) orelse too_general_eqterms(y,x)
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   563
  | too_general_equality _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   564
29267
8615b4f54047 use exists_subterm directly;
wenzelm
parents: 28477
diff changeset
   565
fun has_typed_var tycons = exists_subterm
8615b4f54047 use exists_subterm directly;
wenzelm
parents: 28477
diff changeset
   566
  (fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false);
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   567
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
   568
(* Clauses are forbidden to contain variables of these types. The typical reason
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
   569
   is that they lead to unsoundness. Note that "unit" satisfies numerous
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
   570
   equations like "?x = ()". The resulting clause will have no type constraint,
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
   571
   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
   572
   for higher-order problems. *)
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
   573
val dangerous_types = [@{type_name unit}, @{type_name bool}];
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   574
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
   575
(* Clauses containing variables of type "unit" or "bool" or of the form
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
   576
   "?x = A | ?x = B | ?x = C" are likely to lead to unsound proofs if types are
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
   577
   omitted. *)
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
   578
fun is_dangerous_term _ @{prop True} = true
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
   579
  | is_dangerous_term full_types t =
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   580
    not full_types andalso
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
   581
    (has_typed_var dangerous_types t orelse
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
   582
     forall too_general_equality (HOLogic.disjuncts (strip_Trueprop t)))
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   583
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   584
fun is_fol_goal thy = forall (Meson.is_fol_term thy) o map prop_of
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   585
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
fun relevant_facts full_types respect_no_atp relevance_threshold
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
   587
                   relevance_convergence defs_relevant max_new theory_relevant
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
   588
                   (relevance_override as {add, del, 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
   589
                   (ctxt, (chained_ths, _)) goal_cls =
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   590
  let
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   591
    val thy = ProofContext.theory_of ctxt
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   592
    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
   593
    val has_override = not (null add) orelse not (null del)
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   594
    val is_FO = is_fol_goal thy goal_cls
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   595
    val axioms =
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   596
      checked_name_thm_pairs (respect_no_atp andalso not only) ctxt
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   597
          (map (name_thms_pair_from_ref ctxt chained_ths) add @
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   598
           (if only then []
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   599
            else all_name_thms_pairs respect_no_atp ctxt chained_ths))
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   600
      |> cnf_rules_pairs thy
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   601
      |> not has_override ? make_unique
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   602
      |> filter (fn (cnf_thm, (_, orig_thm)) =>
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   603
                    member Thm.eq_thm add_thms orig_thm orelse
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   604
                    ((not is_FO orelse is_quasi_fol_theorem thy cnf_thm) andalso
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   605
                     not (is_dangerous_term full_types (prop_of cnf_thm))))
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   606
  in
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   607
    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
   608
                     defs_relevant max_new theory_relevant relevance_override
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   609
                     thy axioms (map prop_of goal_cls)
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   610
    |> has_override ? make_unique
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   611
    |> sort (prod_ord string_ord int_ord o pairself (fst o snd))
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   612
  end
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   613
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   614
(** Helper clauses **)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   615
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   616
fun count_combterm (CombConst ((c, _), _, _)) =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   617
    Symtab.map_entry c (Integer.add 1)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   618
  | count_combterm (CombVar _) = I
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   619
  | count_combterm (CombApp (t1, t2)) = count_combterm t1 #> count_combterm t2
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   620
fun count_literal (Literal (_, t)) = count_combterm t
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   621
fun count_clause (HOLClause {literals, ...}) = fold count_literal literals
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   622
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   623
val raw_cnf_rules_pairs = map (fn (name, thm) => (thm, ((name, 0), thm)))
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   624
fun cnf_helper_thms thy raw =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   625
  map (`Thm.get_name_hint)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   626
  #> (if raw then raw_cnf_rules_pairs else cnf_rules_pairs thy)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   627
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   628
val optional_helpers =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   629
  [(["c_COMBI", "c_COMBK"], (false, @{thms COMBI_def COMBK_def})),
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   630
   (["c_COMBB", "c_COMBC"], (false, @{thms COMBB_def COMBC_def})),
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   631
   (["c_COMBS"], (false, @{thms COMBS_def}))]
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   632
val optional_typed_helpers =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   633
  [(["c_True", "c_False"], (true, @{thms True_or_False})),
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   634
   (["c_If"], (true, @{thms if_True if_False True_or_False}))]
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   635
val mandatory_helpers = @{thms fequal_imp_equal equal_imp_fequal}
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   636
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   637
val init_counters =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   638
  Symtab.make (maps (maps (map (rpair 0) o fst))
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   639
                    [optional_helpers, optional_typed_helpers])
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   640
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   641
fun get_helper_clauses thy is_FO full_types conjectures axcls =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   642
  let
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   643
    val axclauses = map snd (make_axiom_clauses thy axcls)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   644
    val ct = fold (fold count_clause) [conjectures, axclauses] init_counters
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   645
    fun is_needed c = the (Symtab.lookup ct c) > 0
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   646
    val cnfs =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   647
      (optional_helpers
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   648
       |> full_types ? append optional_typed_helpers
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   649
       |> maps (fn (ss, (raw, ths)) =>
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   650
                   if exists is_needed ss then cnf_helper_thms thy raw ths
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   651
                   else []))
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   652
      @ (if is_FO then [] else cnf_helper_thms thy false mandatory_helpers)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   653
  in map snd (make_axiom_clauses thy cnfs) end
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   654
31752
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   655
(* prepare for passing to writer,
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   656
   create additional clauses based on the information from extra_cls *)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   657
fun prepare_clauses full_types goal_cls axcls extra_cls thy =
31409
d8537ba165b5 split preparing clauses and writing problemfile;
immler@in.tum.de
parents: 30536
diff changeset
   658
  let
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37359
diff changeset
   659
    val is_FO = is_fol_goal thy goal_cls
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   660
    val ccls = subtract_cls extra_cls goal_cls
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   661
    val _ = app (fn th => trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   662
    val ccltms = map prop_of ccls
31752
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   663
    and axtms = map (prop_of o #1) extra_cls
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   664
    val subs = tfree_classes_of_terms ccltms
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   665
    and supers = tvar_classes_of_terms axtms
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   666
    and tycons = type_consts_of_terms thy (ccltms @ axtms)
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   667
    (*TFrees in conjecture clauses; TVars in axiom clauses*)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   668
    val conjectures = make_conjecture_clauses thy ccls
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   669
    val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses thy extra_cls)
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   670
    val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses thy axcls)
37479
f6b1ee5b420b try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet
parents: 37410
diff changeset
   671
    val helper_clauses =
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   672
      get_helper_clauses thy is_FO full_types conjectures extra_cls
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   673
    val (supers', arity_clauses) = make_arity_clauses thy tycons supers
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   674
    val classrel_clauses = make_classrel_clauses thy subs supers'
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   675
  in
31752
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   676
    (Vector.fromList clnames,
31865
5e97c4abd18e fixed: count constants with supplementary lemmas
immler@in.tum.de
parents: 31837
diff changeset
   677
      (conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses))
31409
d8537ba165b5 split preparing clauses and writing problemfile;
immler@in.tum.de
parents: 30536
diff changeset
   678
  end
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   679
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   680
end;