src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Mon, 17 May 2010 15:21:11 +0200
changeset 36968 62e29faa3718
parent 36922 12f87df9c1a5
child 37171 fc1e20373e6a
permissions -rw-r--r--
make sure chained facts don't pop up in the metis proof
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
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
     8
  type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
     9
  type arity_clause = Sledgehammer_FOL_Clause.arity_clause
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    10
  type axiom_name = Sledgehammer_HOL_Clause.axiom_name
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    11
  type hol_clause = Sledgehammer_HOL_Clause.hol_clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    12
  type hol_clause_id = Sledgehammer_HOL_Clause.hol_clause_id
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    13
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    14
  type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    15
    {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    16
     del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    17
     only: bool}
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    18
22989
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    19
  val tvar_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    20
  val tfree_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    21
  val type_consts_of_terms : theory -> term list -> string list
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
    22
  val get_relevant_facts :
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    23
    bool -> real -> real -> bool -> int -> bool -> relevance_override
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    24
    -> Proof.context * (thm list * 'a) -> thm list
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
    25
    -> (thm * (string * int)) list
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    26
  val prepare_clauses :
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    27
    bool -> thm list -> thm list -> (thm * (axiom_name * hol_clause_id)) list
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    28
    -> (thm * (axiom_name * hol_clause_id)) list -> theory
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    29
    -> axiom_name vector
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    30
       * (hol_clause list * hol_clause list * hol_clause list *
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
    31
          hol_clause list * classrel_clause list * arity_clause list)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    32
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    33
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    34
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    35
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    36
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    37
open Sledgehammer_FOL_Clause
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    38
open Sledgehammer_Fact_Preprocessor
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    39
open Sledgehammer_HOL_Clause
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    40
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    41
type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    42
  {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    43
   del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    44
   only: bool}
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
    45
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    46
(***************************************************************)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    47
(* Relevance Filtering                                         *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    48
(***************************************************************)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    49
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    50
fun strip_Trueprop (@{const Trueprop} $ t) = t
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
    51
  | strip_Trueprop t = t;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    52
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    53
(*A surprising number of theorems contain only a few significant constants.
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    54
  These include all induction rules, and other general theorems. Filtering
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    55
  theorems in clause form reveals these complexities in the form of Skolem 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    56
  functions. If we were instead to filter theorems in their natural form,
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    57
  some other method of measuring theorem complexity would become necessary.*)
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
fun log_weight2 (x:real) = 1.0 + 2.0/Math.ln (x+1.0);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    60
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    61
(*The default seems best in practice. A constant function of one ignores
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    62
  the constant frequencies.*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    63
val weight_fn = log_weight2;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    64
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    65
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    66
(*Including equality in this list might be expected to stop rules like subset_antisym from
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    67
  being chosen, but for some reason filtering works better with them listed. The
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    68
  logical signs All, Ex, &, and --> are omitted because any remaining occurrrences
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    69
  must be within comprehensions.*)
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    70
val standard_consts =
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    71
  [@{const_name Trueprop}, @{const_name "==>"}, @{const_name all},
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    72
   @{const_name "=="}, @{const_name "op |"}, @{const_name Not},
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    73
   @{const_name "op ="}];
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    74
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    75
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    76
(*** constants with types ***)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    77
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    78
(*An abstraction of Isabelle types*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    79
datatype const_typ =  CTVar | CType of string * const_typ list
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    80
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    81
(*Is the second type an instance of the first one?*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    82
fun match_type (CType(con1,args1)) (CType(con2,args2)) = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    83
      con1=con2 andalso match_types args1 args2
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    84
  | match_type CTVar _ = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    85
  | match_type _ CTVar = false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    86
and match_types [] [] = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    87
  | 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
    88
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    89
(*Is there a unifiable constant?*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    90
fun uni_mem gctab (c,c_typ) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    91
  case Symtab.lookup gctab c of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    92
      NONE => false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    93
    | SOME ctyps_list => exists (match_types c_typ) ctyps_list;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    94
  
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    95
(*Maps a "real" type to a const_typ*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    96
fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    97
  | const_typ_of (TFree _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    98
  | const_typ_of (TVar _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    99
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   100
(*Pairs a constant with the list of its type instantiations (using const_typ)*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   101
fun const_with_typ thy (c,typ) = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   102
    let val tvars = Sign.const_typargs thy (c,typ)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   103
    in (c, map const_typ_of tvars) end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   104
    handle TYPE _ => (c,[]);   (*Variable (locale constant): monomorphic*)   
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   105
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   106
(*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
   107
  which we ignore.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   108
fun add_const_typ_table ((c,ctyps), tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   109
  Symtab.map_default (c, [ctyps]) (fn [] => [] | ctyps_list => insert (op =) ctyps ctyps_list) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   110
    tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   111
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   112
(*Free variables are included, as well as constants, to handle locales*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   113
fun add_term_consts_typs_rm thy (Const(c, typ), tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   114
      add_const_typ_table (const_with_typ thy (c,typ), tab) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   115
  | add_term_consts_typs_rm thy (Free(c, typ), tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   116
      add_const_typ_table (const_with_typ thy (c,typ), tab) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   117
  | add_term_consts_typs_rm thy (t $ u, tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   118
      add_term_consts_typs_rm thy (t, add_term_consts_typs_rm thy (u, tab))
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   119
  | add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   120
  | add_term_consts_typs_rm _ (_, tab) = tab;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   121
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   122
(*The empty list here indicates that the constant is being ignored*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   123
fun add_standard_const (s,tab) = Symtab.update (s,[]) tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   124
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   125
val null_const_tab : const_typ list list Symtab.table = 
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   126
    List.foldl add_standard_const Symtab.empty standard_consts;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   127
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30151
diff changeset
   128
fun get_goal_consts_typs thy = List.foldl (add_term_consts_typs_rm thy) null_const_tab;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   129
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   130
(*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
   131
  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
   132
fun const_prop_of theory_relevant th =
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   133
 if theory_relevant then
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   134
  let val name = Context.theory_name (theory_of_thm th)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   135
      val t = Const (name ^ ". 1", HOLogic.boolT)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   136
  in  t $ prop_of th  end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   137
 else prop_of th;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   138
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   139
(**** Constant / Type Frequencies ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   140
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   141
(*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
   142
  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
   143
  a linear ordering on type const_typ list.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   144
  
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   145
local
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   146
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   147
fun cons_nr CTVar = 0
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   148
  | cons_nr (CType _) = 1;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   149
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   150
in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   151
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   152
fun const_typ_ord TU =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   153
  case TU of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   154
    (CType (a, Ts), CType (b, Us)) =>
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   155
      (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
   156
  | (T, U) => int_ord (cons_nr T, cons_nr U);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   157
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   158
end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   159
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31910
diff changeset
   160
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
   161
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   162
fun count_axiom_consts theory_relevant thy ((thm,_), tab) = 
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   163
  let fun count_const (a, T, tab) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   164
        let val (c, cts) = const_with_typ thy (a,T)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   165
        in  (*Two-dimensional table update. Constant maps to types maps to count.*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   166
            Symtab.map_default (c, CTtab.empty) 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   167
                               (CTtab.map_default (cts,0) (fn n => n+1)) tab
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   168
        end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   169
      fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   170
        | count_term_consts (Free(a,T), tab) = count_const(a,T,tab)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   171
        | count_term_consts (t $ u, tab) =
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   172
            count_term_consts (t, count_term_consts (u, tab))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   173
        | count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   174
        | count_term_consts (_, tab) = tab
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   175
  in  count_term_consts (const_prop_of theory_relevant thm, tab)  end;
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
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   178
(**** Actual Filtering Code ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   179
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   180
(*The frequency of a constant is the sum of those of all instances of its type.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   181
fun const_frequency ctab (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
   182
  CTtab.fold (fn (cts', m) => match_types cts cts' ? Integer.add m)
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   183
             (the (Symtab.lookup ctab c)) 0
24287
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
(*Add in a constant's weight, as determined by its frequency.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   186
fun add_ct_weight ctab ((c,T), w) =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   187
  w + weight_fn (real (const_frequency ctab (c,T)));
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   188
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   189
(*Relevant constants are weighted according to frequency, 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   190
  but irrelevant constants are simply counted. Otherwise, Skolem functions,
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   191
  which are rare, would harm a clause's chances of being picked.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   192
fun clause_weight ctab gctyps consts_typs =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   193
    let val rel = filter (uni_mem gctyps) consts_typs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   194
        val rel_weight = List.foldl (add_ct_weight ctab) 0.0 rel
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   195
    in
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   196
        rel_weight / (rel_weight + real (length consts_typs - length rel))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   197
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   198
    
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   199
(*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
   200
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
   201
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   202
fun consts_typs_of_term thy t = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   203
  let val tab = add_term_consts_typs_rm thy (t, null_const_tab)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   204
  in  Symtab.fold add_expand_pairs tab []  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   205
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   206
fun pair_consts_typs_axiom theory_relevant thy (p as (thm, _)) =
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   207
  (p, (consts_typs_of_term thy (const_prop_of theory_relevant thm)));
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   208
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   209
exception ConstFree;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   210
fun dest_ConstFree (Const aT) = aT
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   211
  | dest_ConstFree (Free aT) = aT
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   212
  | dest_ConstFree _ = raise ConstFree;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   213
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   214
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   215
fun defines thy thm gctypes =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   216
    let val tm = prop_of thm
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   217
        fun defs lhs rhs =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   218
            let val (rator,args) = strip_comb lhs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   219
                val ct = const_with_typ thy (dest_ConstFree rator)
33037
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 32994
diff changeset
   220
            in
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 32994
diff changeset
   221
              forall is_Var args andalso uni_mem gctypes ct andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   222
                subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   223
            end
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   224
            handle ConstFree => false
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   225
    in    
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   226
        case tm of
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   227
          @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) => 
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   228
            defs lhs rhs 
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   229
        | _ => false
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   230
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   231
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   232
type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   233
       
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   234
(*For a reverse sort, putting the largest values first.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   235
fun compare_pairs ((_,w1),(_,w2)) = Real.compare (w2,w1);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   236
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   237
(*Limit the number of new clauses, to prevent runaway acceptance.*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   238
fun take_best max_new (newpairs : (annotd_cls*real) list) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   239
  let val nnew = length newpairs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   240
  in
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   241
    if nnew <= max_new then (map #1 newpairs, [])
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   242
    else 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   243
      let val cls = sort compare_pairs newpairs
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   244
          val accepted = List.take (cls, max_new)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   245
      in
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   246
        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
   247
                       ", exceeds the limit of " ^ Int.toString (max_new)));
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   248
        trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   249
        trace_msg (fn () => "Actually passed: " ^
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   250
          space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   251
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   252
        (map #1 accepted, map #1 (List.drop (cls, max_new)))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   253
      end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   254
  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   255
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   256
fun relevant_clauses ctxt relevance_convergence defs_relevant max_new
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   257
                     (relevance_override as {add, del, only}) thy ctab =
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   258
  let
36227
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   259
    val thms_for_facts =
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   260
      maps (maps (cnf_axiom thy) o ProofContext.get_fact ctxt)
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   261
    val add_thms = thms_for_facts add
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   262
    val del_thms = thms_for_facts del
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   263
    fun iter p rel_consts =
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   264
      let
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   265
        fun relevant ([], _) [] = []  (* Nothing added this iteration *)
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   266
          | relevant (newpairs,rejects) [] =
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   267
            let
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   268
              val (newrels, more_rejects) = take_best max_new newpairs
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   269
              val new_consts = maps #2 newrels
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   270
              val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   271
              val newp = p + (1.0 - p) / relevance_convergence
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   272
            in
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   273
              trace_msg (fn () => "relevant this iteration: " ^
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   274
                                  Int.toString (length newrels));
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   275
              map #1 newrels @ iter newp rel_consts' (more_rejects @ rejects)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   276
            end
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   277
          | relevant (newrels, rejects)
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   278
                     ((ax as (clsthm as (thm, (name, n)), consts_typs)) :: axs) =
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   279
            let
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   280
              val weight = if member Thm.eq_thm del_thms thm then 0.0
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   281
                           else if member Thm.eq_thm add_thms thm then 1.0
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   282
                           else if only then 0.0
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   283
                           else clause_weight ctab rel_consts consts_typs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   284
            in
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   285
              if p <= weight orelse
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   286
                 (defs_relevant andalso defines thy (#1 clsthm) rel_consts) then
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   287
                (trace_msg (fn () => name ^ " clause " ^ Int.toString n ^ 
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   288
                                     " passes: " ^ Real.toString weight);
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   289
                relevant ((ax, weight) :: newrels, rejects) axs)
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   290
              else
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   291
                relevant (newrels, ax :: rejects) axs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   292
            end
36182
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   293
        in
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   294
          trace_msg (fn () => "relevant_clauses, current pass mark: " ^
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   295
                              Real.toString p);
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   296
          relevant ([], [])
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   297
        end
b136019c5d61 reorganize Sledgehammer's relevance filter slightly
blanchet
parents: 36061
diff changeset
   298
  in iter end
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   299
        
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   300
fun relevance_filter ctxt relevance_threshold relevance_convergence
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   301
                     defs_relevant max_new theory_relevant relevance_override
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   302
                     thy axioms goals = 
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
   303
  if relevance_threshold > 0.0 then
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   304
    let
36220
f3655a3ae1ab rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents: 36185
diff changeset
   305
      val const_tab = List.foldl (count_axiom_consts theory_relevant thy)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   306
                                 Symtab.empty axioms
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   307
      val goal_const_tab = get_goal_consts_typs thy goals
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   308
      val _ =
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   309
        trace_msg (fn () => "Initial constants: " ^
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   310
                            commas (Symtab.keys goal_const_tab))
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   311
      val relevant =
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   312
        relevant_clauses ctxt relevance_convergence defs_relevant max_new
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   313
                         relevance_override thy const_tab relevance_threshold
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   314
                         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
   315
                         (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
   316
                              axioms)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   317
    in
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   318
      trace_msg (fn () => "Total relevant: " ^ Int.toString (length relevant));
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   319
      relevant
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   320
    end
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   321
  else
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35865
diff changeset
   322
    axioms;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   323
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   324
(***************************************************************)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   325
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   326
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   327
33022
c95102496490 Removal of the unused atpset concept, the atp attribute and some related code.
paulson
parents: 32994
diff changeset
   328
(*** retrieve lemmas and filter them ***)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   329
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   330
(*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
   331
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   332
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
   333
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   334
(*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
   335
  "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
   336
fun is_package_def a =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30291
diff changeset
   337
  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
   338
  in
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   339
     length names > 2 andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   340
     not (hd names = "local") andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   341
     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
   342
  end;
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   343
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   344
fun mk_clause_table xs =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   345
  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
   346
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   347
fun make_unique xs =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   348
  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
   349
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   350
(* Remove existing axiom clauses from the conjecture clauses, as this can
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   351
   dramatically boost an ATP's performance (for some reason). *)
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   352
fun subtract_cls ax_clauses =
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   353
  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
   354
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   355
fun all_valid_thms respect_no_atp ctxt chain_ths =
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   356
  let
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   357
    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
26278
f0c6839df608 replaced obsolete FactIndex.T by Facts.T;
wenzelm
parents: 25761
diff changeset
   358
    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
   359
    val full_space =
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   360
      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
   361
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   362
    fun valid_facts facts =
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   363
      (facts, []) |-> Facts.fold_static (fn (name, ths0) =>
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   364
        let
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   365
          fun check_thms a =
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   366
            (case try (ProofContext.get_thms ctxt) a of
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   367
              NONE => false
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   368
            | SOME ths1 => Thm.eq_thms (ths0, ths1));
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   369
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   370
          val name1 = Facts.extern facts name;
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   371
          val name2 = Name_Space.extern full_space name;
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   372
          val ths = filter_out bad_for_atp ths0;
33641
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   373
        in
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   374
          if Facts.is_concealed facts name orelse
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   375
             forall (member Thm.eq_thm chain_ths) ths orelse
36227
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   376
             (respect_no_atp andalso is_package_def name) then
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   377
            I
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   378
          else case find_first check_thms [name1, name2, name] of
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   379
            NONE => I
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   380
          | SOME a => cons (a, ths)
33641
af07d9cd86ce all_valid_thms: more sophisticated check against global + local name space;
wenzelm
parents: 33316
diff changeset
   381
        end);
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   382
  in valid_facts global_facts @ valid_facts local_facts end;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   383
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   384
fun multi_name a th (n, pairs) =
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   385
  (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   386
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   387
fun add_single_names (a, []) pairs = pairs
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   388
  | add_single_names (a, [th]) pairs = (a, th) :: pairs
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   389
  | add_single_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs));
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   390
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   391
(*Ignore blacklisted basenames*)
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   392
fun add_multi_names (a, ths) pairs =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36550
diff changeset
   393
  if member (op =) multi_base_blacklist (Long_Name.base_name a) then pairs
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   394
  else add_single_names (a, ths) pairs;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   395
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   396
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
   397
36550
f8da913b6c3a make sure short theorem names are preferred to composite ones in Sledgehammer;
blanchet
parents: 36473
diff changeset
   398
(* 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
   399
   names are preferred when both are available. *)
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   400
fun name_thm_pairs respect_no_atp ctxt chain_ths =
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
   401
  let
36058
8256d5a185bd added "respect_no_atp" and "convergence" options to Sledgehammer;
blanchet
parents: 35970
diff changeset
   402
    val (mults, singles) =
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   403
      List.partition is_multi (all_valid_thms respect_no_atp ctxt chain_ths)
36550
f8da913b6c3a make sure short theorem names are preferred to composite ones in Sledgehammer;
blanchet
parents: 36473
diff changeset
   404
    val ps = [] |> fold add_single_names singles
f8da913b6c3a make sure short theorem names are preferred to composite ones in Sledgehammer;
blanchet
parents: 36473
diff changeset
   405
                |> fold add_multi_names mults
36060
4d27652ffb40 reintroduce efficient set structure to collect "no_atp" theorems
blanchet
parents: 36059
diff changeset
   406
  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
   407
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31971
diff changeset
   408
fun check_named ("", th) =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31971
diff changeset
   409
      (warning ("No name for theorem " ^ Display.string_of_thm_without_context th); false)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   410
  | check_named _ = true;
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   411
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   412
fun get_all_lemmas respect_no_atp ctxt chain_ths =
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   413
  let val included_thms =
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   414
        tap (fn ths => trace_msg
33022
c95102496490 Removal of the unused atpset concept, the atp attribute and some related code.
paulson
parents: 32994
diff changeset
   415
                     (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   416
            (name_thm_pairs respect_no_atp ctxt chain_ths)
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   417
  in
31410
c231efe693ce include chain-ths in every prover-call
immler@in.tum.de
parents: 31409
diff changeset
   418
    filter check_named included_thms
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   419
  end;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   420
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   421
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   422
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   423
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   424
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32866
diff changeset
   425
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
   426
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   427
(*Remove this trivial type class*)
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   428
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
   429
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   430
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
   431
  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
   432
  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
   433
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   434
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
   435
  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
   436
  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
   437
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   438
(*fold type constructors*)
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   439
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
   440
  | 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
   441
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   442
val add_type_consts_in_type = fold_type_consts setinsert;
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   443
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   444
(*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
   445
fun add_type_consts_in_term thy =
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   446
  let val const_typargs = Sign.const_typargs thy
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   447
      fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32960
diff changeset
   448
        | add_tcs (Abs (_, _, u)) x = add_tcs u x
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   449
        | add_tcs (t $ u) x = add_tcs t (add_tcs u x)
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   450
        | add_tcs _ x = x
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   451
  in  add_tcs  end
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   452
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   453
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
   454
  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
   455
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   456
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   457
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   458
(* ATP invocation methods setup                                *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   459
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   460
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   461
(*Ensures that no higher-order theorems "leak out"*)
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   462
fun restrict_to_logic thy true cls = filter (Meson.is_fol_term thy o prop_of o fst) cls
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   463
  | restrict_to_logic thy false cls = cls;
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   464
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   465
(**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****)
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   466
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   467
(** 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
   468
  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
   469
  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
   470
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   471
fun occurs ix =
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   472
    let fun occ(Var (jx,_)) = (ix=jx)
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   473
          | occ(t1$t2)      = occ t1 orelse occ t2
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   474
          | occ(Abs(_,_,t)) = occ t
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   475
          | occ _           = false
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   476
    in occ end;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   477
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31410
diff changeset
   478
fun is_recordtype T = not (null (Record.dest_recTs T));
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   479
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   480
(*Unwanted equalities include
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   481
  (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
   482
  (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
   483
*)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   484
fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   485
  | too_general_eqterms _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   486
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   487
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
   488
      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
   489
  | too_general_equality _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   490
29267
8615b4f54047 use exists_subterm directly;
wenzelm
parents: 28477
diff changeset
   491
fun has_typed_var tycons = exists_subterm
8615b4f54047 use exists_subterm directly;
wenzelm
parents: 28477
diff changeset
   492
  (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
   493
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   494
(*Clauses are forbidden to contain variables of these types. The typical reason is that
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   495
  they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=().
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   496
  The resulting clause will have no type constraint, yielding false proofs. Even "bool"
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   497
  leads to many unsound proofs, though (obviously) only for higher-order problems.*)
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   498
val unwanted_types = [@{type_name unit}, @{type_name bool}];
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   499
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   500
fun unwanted t =
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   501
  t = @{prop True} orelse has_typed_var unwanted_types t orelse
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   502
  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
   503
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   504
(*Clauses containing variables of type "unit" or "bool" are unlikely to be useful and
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   505
  likely to lead to unsound proofs.*)
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   506
fun remove_unwanted_clauses cls = filter (not o unwanted o prop_of o fst) cls;
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   507
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
   508
fun is_first_order 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
   509
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   510
fun get_relevant_facts respect_no_atp relevance_threshold relevance_convergence
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   511
                       defs_relevant max_new theory_relevant
36185
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   512
                       (relevance_override as {add, only, ...})
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   513
                       (ctxt, (chain_ths, _)) goal_cls =
36185
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   514
  if (only andalso null add) orelse relevance_threshold > 1.0 then
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   515
    []
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   516
  else
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   517
    let
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   518
      val thy = ProofContext.theory_of ctxt
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
   519
      val is_FO = is_first_order thy goal_cls
36968
62e29faa3718 make sure chained facts don't pop up in the metis proof
blanchet
parents: 36922
diff changeset
   520
      val included_cls = get_all_lemmas respect_no_atp ctxt chain_ths
36185
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   521
        |> cnf_rules_pairs thy |> make_unique
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   522
        |> restrict_to_logic thy is_FO
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   523
        |> remove_unwanted_clauses
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   524
    in
36922
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   525
      relevance_filter ctxt relevance_threshold relevance_convergence
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   526
                       defs_relevant max_new theory_relevant relevance_override
12f87df9c1a5 renamed two Sledgehammer options
blanchet
parents: 36692
diff changeset
   527
                       thy included_cls (map prop_of goal_cls)
36185
0ee736f08ed0 optimize relevance filter by doing a Table.fold directly rather than destroying the datastructure each time;
blanchet
parents: 36182
diff changeset
   528
    end
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   529
31752
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   530
(* prepare for passing to writer,
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   531
   create additional clauses based on the information from extra_cls *)
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
   532
fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy =
31409
d8537ba165b5 split preparing clauses and writing problemfile;
immler@in.tum.de
parents: 30536
diff changeset
   533
  let
32866
1238cbb7c08f atp_minimal using chain_ths again
Philipp Meyer
parents: 32552
diff changeset
   534
    (* add chain thms *)
33306
4138ba02b681 replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents: 33095
diff changeset
   535
    val chain_cls =
36227
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   536
      cnf_rules_pairs thy (filter check_named
8987f7a9afef make Sledgehammer's "add:" and "del:" syntax really work (I hope), by comparing CNF theorems with CNF theorems
blanchet
parents: 36220
diff changeset
   537
                                  (map (`Thm.get_name_hint) chain_ths))
32866
1238cbb7c08f atp_minimal using chain_ths again
Philipp Meyer
parents: 32552
diff changeset
   538
    val axcls = chain_cls @ axcls
1238cbb7c08f atp_minimal using chain_ths again
Philipp Meyer
parents: 32552
diff changeset
   539
    val extra_cls = chain_cls @ extra_cls
36473
8a5c99a1c965 remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
blanchet
parents: 36393
diff changeset
   540
    val is_FO = is_first_order thy goal_cls
36061
d267bdccc660 remove use of Polyhash;
blanchet
parents: 36060
diff changeset
   541
    val ccls = subtract_cls extra_cls goal_cls
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   542
    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
   543
    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
   544
    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
   545
    val subs = tfree_classes_of_terms ccltms
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   546
    and supers = tvar_classes_of_terms axtms
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   547
    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
   548
    (*TFrees in conjecture clauses; TVars in axiom clauses*)
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   549
    val conjectures = make_conjecture_clauses dfg thy ccls
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   550
    val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses dfg thy extra_cls)
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   551
    val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses dfg thy axcls)
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   552
    val helper_clauses = get_helper_clauses dfg thy is_FO (conjectures, extra_cls, [])
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   553
    val (supers', arity_clauses) = make_arity_clauses_dfg dfg thy tycons supers
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   554
    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
   555
  in
31752
19a5f1c8a844 use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents: 31749
diff changeset
   556
    (Vector.fromList clnames,
31865
5e97c4abd18e fixed: count constants with supplementary lemmas
immler@in.tum.de
parents: 31837
diff changeset
   557
      (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
   558
  end
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   559
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   560
end;