src/HOL/Tools/res_atp.ML
author huffman
Thu, 04 Dec 2008 16:05:45 -0800
changeset 28987 dc0ab579a5ca
parent 28477 9339d4dcec8b
child 29267 8615b4f54047
permissions -rw-r--r--
remove duplicated lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
     1
(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     2
    ID: $Id$
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     3
    Copyright 2004 University of Cambridge
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
*)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     5
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     6
signature RES_ATP =
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     7
sig
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
     8
  datatype mode = Auto | Fol | Hol
22989
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
     9
  val tvar_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    10
  val tfree_classes_of_terms : term list -> string list
3bcbe6187027 Added three items to the signature
paulson
parents: 22865
diff changeset
    11
  val type_consts_of_terms : theory -> term list -> string list
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    12
  val write_problem_files : (theory -> bool -> Thm.thm list -> string ->
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    13
  (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list * ResClause.classrelClause list *
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    14
    ResClause.arityClause list -> string list -> ResHolClause.axiom_name list)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    15
    -> int -> bool
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    16
    -> (int -> Path.T) -> Proof.context * thm list * thm
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    17
    -> string list * ResHolClause.axiom_name Vector.vector list
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    18
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    19
22865
da52c2bd66ae renamed call_atp to sledgehammer;
wenzelm
parents: 22824
diff changeset
    20
structure ResAtp: RES_ATP =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    21
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    22
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
    23
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    24
(********************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    25
(* some settings for both background automatic ATP calling procedure*)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    26
(* and also explicit ATP invocation methods                         *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    27
(********************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    28
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    29
(*Translation mode can be auto-detected, or forced to be first-order or higher-order*)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    30
datatype mode = Auto | Fol | Hol;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    31
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    32
val linkup_logic_mode = Auto;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    33
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    34
(*** background linkup ***)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    35
val run_blacklist_filter = true;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
    36
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    37
(*** relevance filter parameters ***)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    38
val run_relevance_filter = true;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    39
val pass_mark = 0.5;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    40
val convergence = 3.2;    (*Higher numbers allow longer inference chains*)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    41
val follow_defs = false;  (*Follow definitions. Makes problems bigger.*)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    42
val include_all = true;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    43
val include_simpset = false;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    44
val include_claset = false;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    45
val include_atpset = true;
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
(***************************************************************)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    48
(* Relevance Filtering                                         *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    49
(***************************************************************)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    50
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
    51
fun strip_Trueprop (Const("Trueprop",_) $ t) = t
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
    52
  | strip_Trueprop t = t;
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    53
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    54
(*A surprising number of theorems contain only a few significant constants.
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    55
  These include all induction rules, and other general theorems. Filtering
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    56
  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
    57
  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
    58
  some other method of measuring theorem complexity would become necessary.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    59
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    60
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
    61
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    62
(*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
    63
  the constant frequencies.*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    64
val weight_fn = log_weight2;
24287
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
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    67
(*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
    68
  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
    69
  logical signs All, Ex, &, and --> are omitted because any remaining occurrrences
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    70
  must be within comprehensions.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    71
val standard_consts = ["Trueprop","==>","all","==","op |","Not","op ="];
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    72
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    73
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    74
(*** constants with types ***)
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
(*An abstraction of Isabelle types*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    77
datatype const_typ =  CTVar | CType of string * const_typ list
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
(*Is the second type an instance of the first one?*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    80
fun match_type (CType(con1,args1)) (CType(con2,args2)) = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    81
      con1=con2 andalso match_types args1 args2
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    82
  | match_type CTVar _ = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    83
  | match_type _ CTVar = false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    84
and match_types [] [] = true
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    85
  | 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
    86
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    87
(*Is there a unifiable constant?*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    88
fun uni_mem gctab (c,c_typ) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    89
  case Symtab.lookup gctab c of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    90
      NONE => false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    91
    | SOME ctyps_list => exists (match_types c_typ) ctyps_list;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    92
  
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    93
(*Maps a "real" type to a const_typ*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    94
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
    95
  | const_typ_of (TFree _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    96
  | const_typ_of (TVar _) = CTVar
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    97
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    98
(*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
    99
fun const_with_typ thy (c,typ) = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   100
    let val tvars = Sign.const_typargs thy (c,typ)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   101
    in (c, map const_typ_of tvars) end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   102
    handle TYPE _ => (c,[]);   (*Variable (locale constant): monomorphic*)   
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   103
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   104
(*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
   105
  which we ignore.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   106
fun add_const_typ_table ((c,ctyps), tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   107
  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
   108
    tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   109
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   110
(*Free variables are included, as well as constants, to handle locales*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   111
fun add_term_consts_typs_rm thy (Const(c, typ), tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   112
      add_const_typ_table (const_with_typ thy (c,typ), tab) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   113
  | add_term_consts_typs_rm thy (Free(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 (t $ u, tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   116
      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
   117
  | add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   118
  | add_term_consts_typs_rm thy (_, tab) = tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   119
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   120
(*The empty list here indicates that the constant is being ignored*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   121
fun add_standard_const (s,tab) = Symtab.update (s,[]) tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   122
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   123
val null_const_tab : const_typ list list Symtab.table = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   124
    foldl add_standard_const Symtab.empty standard_consts;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   125
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   126
fun get_goal_consts_typs thy = foldl (add_term_consts_typs_rm thy) null_const_tab;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   127
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   128
(*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
   129
  takes the given theory into account.*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   130
fun const_prop_of theory_const th =
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   131
 if theory_const then
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   132
  let val name = Context.theory_name (theory_of_thm th)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   133
      val t = Const (name ^ ". 1", HOLogic.boolT)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   134
  in  t $ prop_of th  end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   135
 else prop_of th;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   136
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   137
(**** Constant / Type Frequencies ****)
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
(*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
   140
  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
   141
  a linear ordering on type const_typ list.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   142
  
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   143
local
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
fun cons_nr CTVar = 0
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   146
  | cons_nr (CType _) = 1;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   147
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   148
in
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
fun const_typ_ord TU =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   151
  case TU of
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   152
    (CType (a, Ts), CType (b, Us)) =>
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   153
      (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
   154
  | (T, U) => int_ord (cons_nr T, cons_nr U);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   155
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   156
end;
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
structure CTtab = TableFun(type key = const_typ list val ord = dict_ord const_typ_ord);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   159
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   160
fun count_axiom_consts theory_const thy ((thm,_), tab) = 
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   161
  let fun count_const (a, T, tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   162
	let val (c, cts) = const_with_typ thy (a,T)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   163
	in  (*Two-dimensional table update. Constant maps to types maps to count.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   164
	    Symtab.map_default (c, CTtab.empty) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   165
	                       (CTtab.map_default (cts,0) (fn n => n+1)) tab
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   166
	end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   167
      fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   168
	| count_term_consts (Free(a,T), tab) = count_const(a,T,tab)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   169
	| count_term_consts (t $ u, tab) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   170
	    count_term_consts (t, count_term_consts (u, tab))
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   171
	| count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   172
	| count_term_consts (_, tab) = tab
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   173
  in  count_term_consts (const_prop_of theory_const thm, tab)  end;
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   174
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   175
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   176
(**** Actual Filtering Code ****)
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
(*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
   179
fun const_frequency ctab (c, cts) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   180
  let val pairs = CTtab.dest (the (Symtab.lookup ctab c))
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   181
      fun add ((cts',m), n) = if match_types cts cts' then m+n else n
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   182
  in  List.foldl add 0 pairs  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   183
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   184
(*Add in a constant's weight, as determined by its frequency.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   185
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
   186
  w + weight_fn (real (const_frequency ctab (c,T)));
24287
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
(*Relevant constants are weighted according to frequency, 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   189
  but irrelevant constants are simply counted. Otherwise, Skolem functions,
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   190
  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
   191
fun clause_weight ctab gctyps consts_typs =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   192
    let val rel = filter (uni_mem gctyps) consts_typs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   193
        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
   194
    in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   195
	rel_weight / (rel_weight + real (length consts_typs - length rel))
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
    
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   198
(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   199
fun add_expand_pairs (x,ys) xys = foldl (fn (y,acc) => (x,y)::acc) xys ys;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   200
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   201
fun consts_typs_of_term thy t = 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   202
  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
   203
  in  Symtab.fold add_expand_pairs tab []  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   204
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   205
fun pair_consts_typs_axiom theory_const thy (thm,name) =
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   206
    ((thm,name), (consts_typs_of_term thy (const_prop_of theory_const thm)));
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   207
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   208
exception ConstFree;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   209
fun dest_ConstFree (Const aT) = aT
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   210
  | dest_ConstFree (Free aT) = aT
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   211
  | dest_ConstFree _ = raise ConstFree;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   212
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   213
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   214
fun defines thy (thm,(name,n)) gctypes =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   215
    let val tm = prop_of thm
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   216
	fun defs lhs rhs =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   217
            let val (rator,args) = strip_comb lhs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   218
		val ct = const_with_typ thy (dest_ConstFree rator)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   219
            in  forall is_Var args andalso uni_mem gctypes ct andalso
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   220
                Term.add_vars rhs [] subset Term.add_vars lhs []
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   221
            end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   222
	    handle ConstFree => false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   223
    in    
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   224
	case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   225
		   defs lhs rhs andalso
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   226
		   (Output.debug (fn () => "Definition found: " ^ name ^ "_" ^ Int.toString n); true)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   227
		 | _ => false
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   228
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   229
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   230
type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list);
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
(*For a reverse sort, putting the largest values first.*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   233
fun compare_pairs ((_,w1),(_,w2)) = Real.compare (w2,w1);
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   234
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   235
(*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
   236
fun take_best max_new (newpairs : (annotd_cls*real) list) =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   237
  let val nnew = length newpairs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   238
  in
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   239
    if nnew <= max_new then (map #1 newpairs, [])
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   240
    else 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   241
      let val cls = sort compare_pairs newpairs
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   242
          val accepted = List.take (cls, max_new)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   243
      in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   244
        Output.debug (fn () => ("Number of candidates, " ^ Int.toString nnew ^ 
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   245
		       ", exceeds the limit of " ^ Int.toString (max_new)));
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   246
        Output.debug (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   247
        Output.debug (fn () => "Actually passed: " ^
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   248
          space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   249
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   250
	(map #1 accepted, map #1 (List.drop (cls, max_new)))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   251
      end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   252
  end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   253
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   254
fun relevant_clauses max_new thy ctab p rel_consts =
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   255
  let fun relevant ([],_) [] = [] : (thm * (string * int)) list  (*Nothing added this iteration*)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   256
	| relevant (newpairs,rejects) [] =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   257
	    let val (newrels,more_rejects) = take_best max_new newpairs
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   258
		val new_consts = List.concat (map #2 newrels)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   259
		val rel_consts' = foldl add_const_typ_table rel_consts new_consts
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   260
		val newp = p + (1.0-p) / convergence
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   261
	    in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   262
              Output.debug (fn () => ("relevant this iteration: " ^ Int.toString (length newrels)));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   263
	       (map #1 newrels) @ 
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   264
	       (relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects))
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   265
	    end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   266
	| relevant (newrels,rejects) ((ax as (clsthm as (_,(name,n)),consts_typs)) :: axs) =
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   267
	    let val weight = clause_weight ctab rel_consts consts_typs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   268
	    in
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   269
	      if p <= weight orelse (follow_defs andalso defines thy clsthm rel_consts)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   270
	      then (Output.debug (fn () => (name ^ " clause " ^ Int.toString n ^ 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   271
	                                    " passes: " ^ Real.toString weight));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   272
	            relevant ((ax,weight)::newrels, rejects) axs)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   273
	      else relevant (newrels, ax::rejects) axs
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   274
	    end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   275
    in  Output.debug (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   276
        relevant ([],[]) 
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   277
    end;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   278
	
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   279
fun relevance_filter max_new theory_const thy axioms goals = 
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   280
 if run_relevance_filter andalso pass_mark >= 0.1
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   281
 then
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   282
  let val _ = Output.debug (fn () => "Start of relevance filtering");
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   283
      val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   284
      val goal_const_tab = get_goal_consts_typs thy goals
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   285
      val _ = Output.debug (fn () => ("Initial constants: " ^
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   286
                                 space_implode ", " (Symtab.keys goal_const_tab)));
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   287
      val rels = relevant_clauses max_new thy const_tab (pass_mark) 
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   288
                   goal_const_tab  (map (pair_consts_typs_axiom theory_const thy) axioms)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   289
  in
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   290
      Output.debug (fn () => ("Total relevant: " ^ Int.toString (length rels)));
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   291
      rels
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   292
  end
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   293
 else axioms;
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   294
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   295
(***************************************************************)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   296
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   297
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   298
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   299
(*** white list and black list of lemmas ***)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   300
24215
5458fbf18276 removal of some refs
paulson
parents: 23387
diff changeset
   301
(*The rule subsetI is frequently omitted by the relevance filter. This could be theory data
5458fbf18276 removal of some refs
paulson
parents: 23387
diff changeset
   302
  or identified with ATPset (which however is too big currently)*)
5458fbf18276 removal of some refs
paulson
parents: 23387
diff changeset
   303
val whitelist = [subsetI];
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   304
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   305
(*** retrieve lemmas from clasimpset and atpset, may filter them ***)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   306
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   307
(*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
   308
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   309
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
   310
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   311
(*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
   312
  "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
   313
fun is_package_def a =
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21790
diff changeset
   314
  let val names = NameSpace.explode a
21690
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   315
  in
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   316
     length names > 2 andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   317
     not (hd names = "local") andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   318
     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
   319
  end;
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   320
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   321
(** a hash function from Term.term to int, and also a hash table **)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   322
val xor_words = List.foldl Word.xorb 0w0;
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   323
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   324
fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w)
20661
46832fee1215 Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
paulson
parents: 20643
diff changeset
   325
  | hashw_term ((Free(a,_)), w) = Polyhash.hashw_string (a,w)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   326
  | hashw_term ((Var(_,_)), w) = w
20661
46832fee1215 Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
paulson
parents: 20643
diff changeset
   327
  | hashw_term ((Bound i), w) = Polyhash.hashw_int(i,w)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   328
  | hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   329
  | hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w)));
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   330
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
   331
fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_term(P,0w0))
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
   332
  | hash_literal P = hashw_term(P,0w0);
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   333
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   334
fun hash_term t = Word.toIntX (xor_words (map hash_literal (HOLogic.disjuncts (strip_Trueprop t))));
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   335
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   336
fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2);
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   337
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   338
exception HASH_CLAUSE;
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   339
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   340
(*Create a hash table for clauses, of the given size*)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   341
fun mk_clause_table n =
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   342
      Polyhash.mkTable (hash_term o prop_of, equal_thm)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   343
                       (n, HASH_CLAUSE);
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   344
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   345
(*Use a hash table to eliminate duplicates from xs. Argument is a list of
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   346
  (thm * (string * int)) tuples. The theorems are hashed into the table. *)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   347
fun make_unique xs =
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   348
  let val ht = mk_clause_table 7000
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   349
  in
22130
0906fd95e0b5 Output.debug: non-strict;
wenzelm
parents: 22012
diff changeset
   350
      Output.debug (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   351
      app (ignore o Polyhash.peekInsert ht) xs;
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   352
      Polyhash.listItems ht
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   353
  end;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   354
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   355
(*Remove existing axiom clauses from the conjecture clauses, as this can dramatically
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   356
  boost an ATP's performance (for some reason)*)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   357
fun subtract_cls c_clauses ax_clauses =
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   358
  let val ht = mk_clause_table 2200
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   359
      fun known x = isSome (Polyhash.peek ht x)
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   360
  in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   361
      app (ignore o Polyhash.peekInsert ht) ax_clauses;
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   362
      filter (not o known) c_clauses
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   363
  end;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   364
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   365
fun valid_facts facts =
26691
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   366
  Facts.fold_static (fn (name, ths) =>
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   367
    if run_blacklist_filter andalso is_package_def name then I
26691
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   368
    else
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   369
      let val xname = Facts.extern facts name in
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   370
        if NameSpace.is_hidden xname then I
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   371
        else cons (xname, filter_out ResAxioms.bad_for_atp ths)
520c99e0b9a0 valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents: 26675
diff changeset
   372
      end) facts [];
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   373
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   374
fun all_valid_thms ctxt =
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   375
  let
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   376
    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
26278
f0c6839df608 replaced obsolete FactIndex.T by Facts.T;
wenzelm
parents: 25761
diff changeset
   377
    val local_facts = ProofContext.facts_of ctxt;
26675
fba93ce71435 all_valid_thms: use new facts tables;
wenzelm
parents: 26662
diff changeset
   378
  in valid_facts global_facts @ valid_facts local_facts end;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   379
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   380
fun multi_name a (th, (n,pairs)) =
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   381
  (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs)
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   382
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   383
fun add_single_names ((a, []), pairs) = pairs
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   384
  | add_single_names ((a, [th]), pairs) = (a,th)::pairs
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   385
  | add_single_names ((a, ths), pairs) = #2 (foldl (multi_name a) (1,pairs) ths);
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   386
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   387
(*Ignore blacklisted basenames*)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   388
fun add_multi_names ((a, ths), pairs) =
24854
0ebcd575d3c6 filtering out some package theorems
paulson
parents: 24827
diff changeset
   389
  if (Sign.base_name a) mem_string ResAxioms.multi_base_blacklist  then pairs
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   390
  else add_single_names ((a, ths), pairs);
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   391
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   392
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
   393
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   394
(*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   395
fun name_thm_pairs ctxt =
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   396
  let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt)
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   397
      val ht = mk_clause_table 900   (*ht for blacklisted theorems*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   398
      fun blacklisted x = run_blacklist_filter andalso isSome (Polyhash.peek ht x)
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   399
  in
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   400
      app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt);
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   401
      filter (not o blacklisted o #2)
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   402
        (foldl add_single_names (foldl add_multi_names [] mults) singles)
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24215
diff changeset
   403
  end;
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   404
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 26691
diff changeset
   405
fun check_named ("",th) = (warning ("No name for theorem " ^ Display.string_of_thm th); false)
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   406
  | check_named (_,th) = true;
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   407
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 26691
diff changeset
   408
fun display_thm (name,th) = Output.debug (fn () => name ^ ": " ^ Display.string_of_thm th);
22193
62753ae847a2 Streamlined and improved debugging messages
paulson
parents: 22130
diff changeset
   409
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   410
(* get lemmas from claset, simpset, atpset and extra supplied rules *)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   411
fun get_clasimp_atp_lemmas ctxt user_thms =
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   412
  let val included_thms =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   413
        if include_all
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   414
        then (tap (fn ths => Output.debug
22130
0906fd95e0b5 Output.debug: non-strict;
wenzelm
parents: 22012
diff changeset
   415
                     (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   416
                  (name_thm_pairs ctxt))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   417
        else
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   418
        let val claset_thms =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   419
                if include_claset then ResAxioms.claset_rules_of ctxt
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   420
                else []
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   421
            val simpset_thms =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   422
                if include_simpset then ResAxioms.simpset_rules_of ctxt
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   423
                else []
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   424
            val atpset_thms =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   425
                if include_atpset then ResAxioms.atpset_rules_of ctxt
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   426
                else []
22193
62753ae847a2 Streamlined and improved debugging messages
paulson
parents: 22130
diff changeset
   427
            val _ = (Output.debug (fn () => "ATP theorems: ");  app display_thm atpset_thms)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   428
        in  claset_thms @ simpset_thms @ atpset_thms  end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   429
      val user_rules = filter check_named
22382
dbf09db0a40d New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents: 22217
diff changeset
   430
                         (map ResAxioms.pairname
24215
5458fbf18276 removal of some refs
paulson
parents: 23387
diff changeset
   431
                           (if null user_thms then whitelist else user_thms))
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   432
  in
21224
f86b8463af37 Proper theorem names at last, no fakes!!
paulson
parents: 21132
diff changeset
   433
      (filter check_named included_thms, user_rules)
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19768
diff changeset
   434
  end;
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   435
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   436
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   437
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   438
(***************************************************************)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   439
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   440
fun add_classes (sorts, cset) = foldl setinsert cset (List.concat sorts);
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   441
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   442
(*Remove this trivial type class*)
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   443
fun delete_type cset = Symtab.delete_safe "HOL.type" cset;
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   444
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   445
fun tvar_classes_of_terms ts =
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   446
  let val sorts_list = map (map #2 o term_tvars) ts
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   447
  in  Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list))  end;
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   448
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   449
fun tfree_classes_of_terms ts =
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   450
  let val sorts_list = map (map #2 o term_tfrees) ts
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   451
  in  Symtab.keys (delete_type (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
   452
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   453
(*fold type constructors*)
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   454
fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   455
  | fold_type_consts f T x = x;
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   456
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   457
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
   458
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   459
(*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
   460
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
   461
  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
   462
      fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   463
        | add_tcs (Abs (_, T, u)) x = add_tcs u x
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   464
        | 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
   465
        | add_tcs _ x = x
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   466
  in  add_tcs  end
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   467
21397
2134b81a0b37 Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents: 21373
diff changeset
   468
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
   469
  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
   470
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   471
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   472
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   473
(* ATP invocation methods setup                                *)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   474
(***************************************************************)
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   475
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   476
(*Ensures that no higher-order theorems "leak out"*)
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   477
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
   478
  | 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
   479
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   480
(**** 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
   481
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   482
(** 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
   483
  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
   484
  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
   485
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   486
fun occurs ix =
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   487
    let fun occ(Var (jx,_)) = (ix=jx)
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   488
          | occ(t1$t2)      = occ t1 orelse occ t2
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   489
          | occ(Abs(_,_,t)) = occ t
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   490
          | occ _           = false
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   491
    in occ end;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   492
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   493
fun is_recordtype T = not (null (RecordPackage.dest_recTs T));
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   494
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   495
(*Unwanted equalities include
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   496
  (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
   497
  (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
   498
*)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   499
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
   500
  | too_general_eqterms _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   501
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   502
fun too_general_equality (Const ("op =", _) $ x $ y) =
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   503
      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
   504
  | too_general_equality _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   505
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   506
(* tautologous? *)
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   507
fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   508
  | is_taut _ = false;
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   509
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   510
(*True if the term contains a variable whose (atomic) type is in the given list.*)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   511
fun has_typed_var tycons =
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   512
  let fun var_tycon (Var (_, Type(a,_))) = a mem_string tycons
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   513
        | var_tycon _ = false
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   514
  in  exists var_tycon o term_vars  end;
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   515
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   516
(*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
   517
  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
   518
  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
   519
  leads to many unsound proofs, though (obviously) only for higher-order problems.*)
24215
5458fbf18276 removal of some refs
paulson
parents: 23387
diff changeset
   520
val unwanted_types = ["Product_Type.unit","bool"];
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   521
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   522
fun unwanted t =
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   523
  is_taut t orelse has_typed_var unwanted_types t orelse
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   524
  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
   525
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   526
(*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
   527
  likely to lead to unsound proofs.*)
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   528
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
   529
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   530
(* TODO: problem file for *one* subgoal would be sufficient *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   531
(*Write out problem files for each subgoal.
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   532
  Argument probfile generates filenames from subgoal-number
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   533
  Argument writer is either a tptp_write_file or dfg_write_file from ResHolClause
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   534
  Arguments max_new and theory_const are booleans controlling relevance_filter
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   535
    (necessary for different provers)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   536
  *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   537
fun write_problem_files writer max_new theory_const probfile (ctxt, chain_ths, th) =
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   538
  let val goals = Thm.prems_of th
17717
7c6a96cbc966 improvements for problem generation
paulson
parents: 17690
diff changeset
   539
      val thy = ProofContext.theory_of ctxt
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20446
diff changeset
   540
      fun get_neg_subgoals [] _ = []
22865
da52c2bd66ae renamed call_atp to sledgehammer;
wenzelm
parents: 22824
diff changeset
   541
        | get_neg_subgoals (gl::gls) n = #1 (ResAxioms.neg_conjecture_clauses th n) ::
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21980
diff changeset
   542
                                         get_neg_subgoals gls (n+1)
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   543
      val goal_cls = get_neg_subgoals goals 1
25243
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25085
diff changeset
   544
                     handle THM ("assume: variables", _, _) => 
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25085
diff changeset
   545
                       error "Sledgehammer: Goal contains type variables (TVars)"                       
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   546
      val isFO = case linkup_logic_mode of
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   547
			  Auto => forall (Meson.is_fol_term thy) (map prop_of (List.concat goal_cls))
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   548
			| Fol => true
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   549
			| Hol => false
24546
c90cee3163b7 chained facts are now included
paulson
parents: 24425
diff changeset
   550
      val (included_thms,white_thms) = get_clasimp_atp_lemmas ctxt chain_ths
27178
c8ddb3000743 ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 26928
diff changeset
   551
      val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   552
                                       |> restrict_to_logic thy isFO
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   553
                                       |> remove_unwanted_clauses
27178
c8ddb3000743 ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 26928
diff changeset
   554
      val white_cls = ResAxioms.cnf_rules_pairs thy white_thms
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   555
      (*clauses relevant to goal gl*)
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   556
      val axcls_list = map (fn ngcls => white_cls @ relevance_filter max_new theory_const thy included_cls (map prop_of ngcls)) goal_cls
20526
756c4f1fd04c Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents: 20457
diff changeset
   557
      fun write_all [] [] _ = []
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21563
diff changeset
   558
        | write_all (ccls::ccls_list) (axcls::axcls_list) k =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
   559
            let val fname = File.platform_path (probfile k)
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   560
                val axcls = make_unique axcls
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 26691
diff changeset
   561
                val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls
20868
724ab9896068 Improved detection of identical clauses, also in the conjecture
paulson
parents: 20854
diff changeset
   562
                val ccls = subtract_cls ccls axcls
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 26691
diff changeset
   563
                val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   564
                val ccltms = map prop_of ccls
21373
18f519614978 Arity clauses are now produced only for types and type classes actually used.
paulson
parents: 21311
diff changeset
   565
                and axtms = map (prop_of o #1) axcls
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   566
                val subs = tfree_classes_of_terms ccltms
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   567
                and supers = tvar_classes_of_terms axtms
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   568
                and tycons = type_consts_of_terms thy (ccltms@axtms)
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21243
diff changeset
   569
                (*TFrees in conjecture clauses; TVars in axiom clauses*)
22645
8a70bc644833 Improved treatment of classrel/arity clauses
paulson
parents: 22578
diff changeset
   570
                val (supers',arity_clauses) = ResClause.make_arity_clauses thy tycons supers
8a70bc644833 Improved treatment of classrel/arity clauses
paulson
parents: 22578
diff changeset
   571
                val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24920
diff changeset
   572
                val clnames = writer thy isFO ccls fname (axcls,classrel_clauses,arity_clauses) []
21888
c75a44597fb7 change from "Array" to "Vector"
paulson
parents: 21858
diff changeset
   573
                val thm_names = Vector.fromList clnames
20870
992bcbff055a Tidied code to delete tmp files.
paulson
parents: 20868
diff changeset
   574
            in  (thm_names,fname) :: write_all ccls_list axcls_list (k+1)  end
992bcbff055a Tidied code to delete tmp files.
paulson
parents: 20868
diff changeset
   575
      val (thm_names_list, filenames) = ListPair.unzip (write_all goal_cls axcls_list 1)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   576
  in
20870
992bcbff055a Tidied code to delete tmp files.
paulson
parents: 20868
diff changeset
   577
      (filenames, thm_names_list)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
   578
  end;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   579
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   580
end;